| author | wenzelm | 
| Sat, 30 Sep 2017 20:06:26 +0200 | |
| changeset 66732 | e566fb4d43d4 | 
| parent 66330 | dcb3e6bdc00a | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 31048 
ac146fc38b51
refined HOL string theories and corresponding ML fragments
 haftmann parents: 
31031diff
changeset | 1 | (* Author: Florian Haftmann, TU Muenchen *) | 
| 26168 | 2 | |
| 60758 | 3 | section \<open>Reflecting Pure types into HOL\<close> | 
| 26168 | 4 | |
| 28952 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
28562diff
changeset | 5 | theory Typerep | 
| 51112 | 6 | imports String | 
| 26168 | 7 | begin | 
| 8 | ||
| 58310 | 9 | datatype typerep = Typerep String.literal "typerep list" | 
| 26168 | 10 | |
| 28335 | 11 | class typerep = | 
| 31031 | 12 | fixes typerep :: "'a itself \<Rightarrow> typerep" | 
| 26168 | 13 | begin | 
| 14 | ||
| 29574 | 15 | definition typerep_of :: "'a \<Rightarrow> typerep" where | 
| 28335 | 16 |   [simp]: "typerep_of x = typerep TYPE('a)"
 | 
| 26168 | 17 | |
| 18 | end | |
| 19 | ||
| 35115 | 20 | syntax | 
| 21 |   "_TYPEREP" :: "type => logic"  ("(1TYPEREP/(1'(_')))")
 | |
| 22 | ||
| 60758 | 23 | parse_translation \<open> | 
| 52143 | 24 | let | 
| 25 | fun typerep_tr (*"_TYPEREP"*) [ty] = | |
| 26 |           Syntax.const @{const_syntax typerep} $
 | |
| 56243 | 27 |             (Syntax.const @{syntax_const "_constrain"} $ Syntax.const @{const_syntax Pure.type} $
 | 
| 52143 | 28 |               (Syntax.const @{type_syntax itself} $ ty))
 | 
| 29 |       | typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts);
 | |
| 30 |   in [(@{syntax_const "_TYPEREP"}, K typerep_tr)] end
 | |
| 60758 | 31 | \<close> | 
| 35115 | 32 | |
| 60758 | 33 | typed_print_translation \<open> | 
| 52143 | 34 | let | 
| 35 | fun typerep_tr' ctxt (*"typerep"*) | |
| 36 |             (Type (@{type_name fun}, [Type (@{type_name itself}, [T]), _]))
 | |
| 56243 | 37 |             (Const (@{const_syntax Pure.type}, _) :: ts) =
 | 
| 52143 | 38 | Term.list_comb | 
| 39 |             (Syntax.const @{syntax_const "_TYPEREP"} $ Syntax_Phases.term_of_typ ctxt T, ts)
 | |
| 40 | | typerep_tr' _ T ts = raise Match; | |
| 41 |   in [(@{const_syntax typerep}, typerep_tr')] end
 | |
| 60758 | 42 | \<close> | 
| 26168 | 43 | |
| 60758 | 44 | setup \<open> | 
| 31137 | 45 | let | 
| 26168 | 46 | |
| 31137 | 47 | fun add_typerep tyco thy = | 
| 26168 | 48 | let | 
| 28335 | 49 |     val sorts = replicate (Sign.arity_number thy tyco) @{sort typerep};
 | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
42247diff
changeset | 50 | val vs = Name.invent_names Name.context "'a" sorts; | 
| 26168 | 51 | val ty = Type (tyco, map TFree vs); | 
| 28335 | 52 |     val lhs = Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep})
 | 
| 26168 | 53 |       $ Free ("T", Term.itselfT ty);
 | 
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
31137diff
changeset | 54 |     val rhs = @{term Typerep} $ HOLogic.mk_literal tyco
 | 
| 31137 | 55 |       $ HOLogic.mk_list @{typ typerep} (map (HOLogic.mk_typerep o TFree) vs);
 | 
| 26168 | 56 | val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); | 
| 57 | in | |
| 58 | thy | |
| 38348 
cf7b2121ad9d
moved instantiation target formally to class_target.ML
 haftmann parents: 
36176diff
changeset | 59 |     |> Class.instantiation ([tyco], vs, @{sort typerep})
 | 
| 26168 | 60 | |> `(fn lthy => Syntax.check_term lthy eq) | 
| 63352 | 61 | |-> (fn eq => Specification.definition NONE [] [] (Binding.empty_atts, eq)) | 
| 26168 | 62 | |> snd | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58889diff
changeset | 63 | |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt []) | 
| 26168 | 64 | end; | 
| 65 | ||
| 45693 | 66 | fun ensure_typerep tyco thy = | 
| 48272 | 67 |   if not (Sorts.has_instance (Sign.classes_of thy) tyco @{sort typerep})
 | 
| 68 |     andalso Sorts.has_instance (Sign.classes_of thy) tyco @{sort type}
 | |
| 31137 | 69 | then add_typerep tyco thy else thy; | 
| 70 | ||
| 71 | in | |
| 26168 | 72 | |
| 31137 | 73 | add_typerep @{type_name fun}
 | 
| 58662 | 74 | #> Typedef.interpretation (Local_Theory.background_theory o ensure_typerep) | 
| 66330 
dcb3e6bdc00a
one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
 haftmann parents: 
63352diff
changeset | 75 | #> Code.type_interpretation ensure_typerep | 
| 26168 | 76 | |
| 31137 | 77 | end | 
| 60758 | 78 | \<close> | 
| 26168 | 79 | |
| 28562 | 80 | lemma [code]: | 
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38348diff
changeset | 81 | "HOL.equal (Typerep tyco1 tys1) (Typerep tyco2 tys2) \<longleftrightarrow> HOL.equal tyco1 tyco2 | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38348diff
changeset | 82 | \<and> list_all2 HOL.equal tys1 tys2" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38348diff
changeset | 83 | by (auto simp add: eq_equal [symmetric] list_all2_eq [symmetric]) | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38348diff
changeset | 84 | |
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38348diff
changeset | 85 | lemma [code nbe]: | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38348diff
changeset | 86 | "HOL.equal (x :: typerep) x \<longleftrightarrow> True" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38348diff
changeset | 87 | by (fact equal_refl) | 
| 26168 | 88 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 89 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 90 | type_constructor typerep \<rightharpoonup> (Eval) "Term.typ" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 91 | | constant Typerep \<rightharpoonup> (Eval) "Term.Type/ (_, _)" | 
| 26168 | 92 | |
| 31031 | 93 | code_reserved Eval Term | 
| 26168 | 94 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35430diff
changeset | 95 | hide_const (open) typerep Typerep | 
| 26168 | 96 | |
| 97 | end |