aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorJosh Chen2020-06-03 13:09:30 +0200
committerJosh Chen2020-06-03 13:09:30 +0200
commite513fc2958133e3a00c06ebcd1214741843acf08 (patch)
tree717ab2e413e8fee071d1e4b22b2ed224f473d081
parent8c0205ade477f732fceb484bf5986fdbc6469667 (diff)
1. Type information context data
2. Small reformulations of rules 3. Bool
-rw-r--r--spartan/core/Spartan.thy5
-rw-r--r--spartan/core/lib/types.ML18
-rw-r--r--spartan/data/List.thy6
-rw-r--r--spartan/data/More_Types.thy8
4 files changed, 32 insertions, 5 deletions
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 \<Longrightarrow> A: U (S i)"
by (erule U_cumulative, rule lt_S)
@@ -154,6 +154,7 @@ named_theorems typechk
ML_file \<open>lib/lib.ML\<close>
ML_file \<open>lib/goals.ML\<close>
ML_file \<open>lib/focus.ML\<close>
+ML_file \<open>lib/types.ML\<close>
section \<open>Congruence automation\<close>
@@ -322,7 +323,7 @@ lemma lift_universe_codomain:
shows "f: A \<rightarrow> 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 \<open>Function composition\<close>
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 \<equiv> ListInd A (\<lambda>_. C)"
Lemma (derive) ListCase:
assumes
- "A: U i" "\<And>xs. xs: List A \<Longrightarrow> C xs: U i" and
+ "xs: List A" and
nil_case: "c\<^sub>0: C (nil A)" and
cons_case: "\<And>x xs. \<lbrakk>x: A; xs: List A\<rbrakk> \<Longrightarrow> f x xs: C (cons A x xs)" and
- "xs: List A"
+ "\<And>xs. xs: List A \<Longrightarrow> C xs: U i"
shows "?List_cases A (\<lambda>xs. C xs) c\<^sub>0 (\<lambda>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 \<open>Notation\<close>
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 \<equiv> \<top> \<or> \<top>"
definition "true \<equiv> inl \<top> \<top> tt"
definition "false \<equiv> inr \<top> \<top> 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
+
\<comment> \<open>Can define if-then-else etc.\<close>