author | wenzelm |
Fri, 29 Nov 2024 17:40:15 +0100 | |
changeset 81507 | 08574da77b4a |
parent 81125 | ec121999a9cb |
child 81706 | 7beb0cf38292 |
permissions | -rw-r--r-- |
31048
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
31031
diff
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:
28562
diff
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 |
81125 | 21 |
"_TYPEREP" :: "type => logic" (\<open>(\<open>indent=1 notation=\<open>mixfix TYPEREP\<close>\<close>TYPEREP/(1'(_')))\<close>) |
80760 | 22 |
syntax_consts |
23 |
"_TYPEREP" \<rightleftharpoons> typerep |
|
35115 | 24 |
|
60758 | 25 |
parse_translation \<open> |
52143 | 26 |
let |
27 |
fun typerep_tr (*"_TYPEREP"*) [ty] = |
|
69593 | 28 |
Syntax.const \<^const_syntax>\<open>typerep\<close> $ |
29 |
(Syntax.const \<^syntax_const>\<open>_constrain\<close> $ Syntax.const \<^const_syntax>\<open>Pure.type\<close> $ |
|
30 |
(Syntax.const \<^type_syntax>\<open>itself\<close> $ ty)) |
|
52143 | 31 |
| typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts); |
69593 | 32 |
in [(\<^syntax_const>\<open>_TYPEREP\<close>, K typerep_tr)] end |
60758 | 33 |
\<close> |
35115 | 34 |
|
60758 | 35 |
typed_print_translation \<open> |
52143 | 36 |
let |
74375 | 37 |
fun typerep_tr' ctxt (*"typerep"*) \<^Type>\<open>fun \<^Type>\<open>itself T\<close> _\<close> |
69593 | 38 |
(Const (\<^const_syntax>\<open>Pure.type\<close>, _) :: ts) = |
52143 | 39 |
Term.list_comb |
69593 | 40 |
(Syntax.const \<^syntax_const>\<open>_TYPEREP\<close> $ Syntax_Phases.term_of_typ ctxt T, ts) |
52143 | 41 |
| typerep_tr' _ T ts = raise Match; |
69593 | 42 |
in [(\<^const_syntax>\<open>typerep\<close>, typerep_tr')] end |
60758 | 43 |
\<close> |
26168 | 44 |
|
60758 | 45 |
setup \<open> |
31137 | 46 |
let |
26168 | 47 |
|
31137 | 48 |
fun add_typerep tyco thy = |
26168 | 49 |
let |
69593 | 50 |
val sorts = replicate (Sign.arity_number thy tyco) \<^sort>\<open>typerep\<close>; |
81507 | 51 |
val vs = Name.invent_types_global sorts; |
26168 | 52 |
val ty = Type (tyco, map TFree vs); |
74345 | 53 |
val lhs = \<^Const>\<open>typerep ty\<close> $ Free ("T", Term.itselfT ty); |
54 |
val rhs = \<^Const>\<open>Typerep\<close> $ HOLogic.mk_literal tyco |
|
55 |
$ HOLogic.mk_list \<^Type>\<open>typerep\<close> (map (HOLogic.mk_typerep o TFree) vs); |
|
26168 | 56 |
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); |
57 |
in |
|
58 |
thy |
|
69593 | 59 |
|> Class.instantiation ([tyco], vs, \<^sort>\<open>typerep\<close>) |
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:
58889
diff
changeset
|
63 |
|> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt []) |
26168 | 64 |
end; |
65 |
||
45693 | 66 |
fun ensure_typerep tyco thy = |
69593 | 67 |
if not (Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\<open>typerep\<close>) |
68 |
andalso Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\<open>type\<close> |
|
31137 | 69 |
then add_typerep tyco thy else thy; |
70 |
||
71 |
in |
|
26168 | 72 |
|
69593 | 73 |
add_typerep \<^type_name>\<open>fun\<close> |
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:
63352
diff
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:
38348
diff
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:
38348
diff
changeset
|
82 |
\<and> list_all2 HOL.equal tys1 tys2" |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38348
diff
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:
38348
diff
changeset
|
84 |
|
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38348
diff
changeset
|
85 |
lemma [code nbe]: |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38348
diff
changeset
|
86 |
"HOL.equal (x :: typerep) x \<longleftrightarrow> True" |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38348
diff
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:
52143
diff
changeset
|
89 |
code_printing |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52143
diff
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:
52143
diff
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:
35430
diff
changeset
|
95 |
hide_const (open) typerep Typerep |
26168 | 96 |
|
97 |
end |