From e513fc2958133e3a00c06ebcd1214741843acf08 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Wed, 3 Jun 2020 13:09:30 +0200 Subject: 1. Type information context data 2. Small reformulations of rules 3. Bool --- spartan/core/Spartan.thy | 5 +++-- spartan/core/lib/types.ML | 18 ++++++++++++++++++ spartan/data/List.thy | 6 +++--- spartan/data/More_Types.thy | 8 ++++++++ 4 files changed, 32 insertions(+), 5 deletions(-) create mode 100644 spartan/core/lib/types.ML diff --git a/spartan/core/Spartan.thy b/spartan/core/Spartan.thy index 9a6d884..55eb657 100644 --- a/spartan/core/Spartan.thy +++ b/spartan/core/Spartan.thy @@ -56,7 +56,7 @@ lemma U_in_U: "U i: U (S i)" by (rule U_hierarchy, rule lt_S) -lemma lift_universe: +lemma lift_U: "A: U i \ A: U (S i)" by (erule U_cumulative, rule lt_S) @@ -154,6 +154,7 @@ named_theorems typechk ML_file \lib/lib.ML\ ML_file \lib/goals.ML\ ML_file \lib/focus.ML\ +ML_file \lib/types.ML\ section \Congruence automation\ @@ -322,7 +323,7 @@ lemma lift_universe_codomain: shows "f: A \ U (S j)" apply (sub eta_exp) apply known - apply (Pure.rule intros; rule lift_universe) + apply (Pure.rule intros; rule lift_U) done subsection \Function composition\ diff --git a/spartan/core/lib/types.ML b/spartan/core/lib/types.ML new file mode 100644 index 0000000..b0792fe --- /dev/null +++ b/spartan/core/lib/types.ML @@ -0,0 +1,18 @@ +structure Types += struct + +structure Data = Generic_Data ( + type T = thm Item_Net.T + val empty = Item_Net.init Thm.eq_thm + (single o Lib.term_of_typing o Thm.prop_of) + val extend = I + val merge = Item_Net.merge +) + +fun put_type typing = Context.proof_map (Data.map (Item_Net.update typing)) +fun put_types typings = foldr1 (op o) (map put_type typings) + +fun get_types ctxt tm = Item_Net.retrieve (Data.get (Context.Proof ctxt)) tm + + +end diff --git a/spartan/data/List.thy b/spartan/data/List.thy index 8fdaa1d..1798a23 100644 --- a/spartan/data/List.thy +++ b/spartan/data/List.thy @@ -52,14 +52,14 @@ abbreviation "ListRec A C \ ListInd A (\_. C)" Lemma (derive) ListCase: assumes - "A: U i" "\xs. xs: List A \ C xs: U i" and + "xs: List A" and nil_case: "c\<^sub>0: C (nil A)" and cons_case: "\x xs. \x: A; xs: List A\ \ f x xs: C (cons A x xs)" and - "xs: List A" + "\xs. xs: List A \ C xs: U i" shows "?List_cases A (\xs. C xs) c\<^sub>0 (\x xs. f x xs) xs: C xs" by (elim xs) (fact nil_case, rule cons_case) -lemmas List_cases [cases] = ListCase[rotated 4] +lemmas List_cases [cases] = ListCase[unfolded ListCase_def] section \Notation\ diff --git a/spartan/data/More_Types.thy b/spartan/data/More_Types.thy index 625f639..1d7abb9 100644 --- a/spartan/data/More_Types.thy +++ b/spartan/data/More_Types.thy @@ -90,6 +90,14 @@ definition "Bool \ \ \ \" definition "true \ inl \ \ tt" definition "false \ inr \ \ tt" +Lemma + BoolF: "Bool: U i" and + Bool_true: "true: Bool" and + Bool_false: "false: Bool" + unfolding Bool_def true_def false_def by typechk+ + +lemmas [intros] = BoolF Bool_true Bool_false + \ \Can define if-then-else etc.\ -- cgit v1.2.3