diff options
author | Josh Chen | 2020-06-03 13:10:35 +0200 |
---|---|---|
committer | Josh Chen | 2020-06-03 13:10:35 +0200 |
commit | 9050b7414021db31b23a034567ebc6da3f6c5f67 (patch) | |
tree | 2f5928e275c3e8ae7d99fef65e7fddf31f11d2f6 /spartan/core | |
parent | ccc26cf8073071698f333107fd7443a6af7fb3de (diff) | |
parent | 515c142828e66dcb1c273e53816ef8b6e1bb3f01 (diff) |
Merge branch 'dev'
Diffstat (limited to 'spartan/core')
-rw-r--r-- | spartan/core/Spartan.thy | 5 | ||||
-rw-r--r-- | spartan/core/lib/types.ML | 18 |
2 files changed, 21 insertions, 2 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 |