1(* Author: Florian Haftmann, TU Muenchen *)
2
3section \<open>Reflecting Pure types into HOL\<close>
4
5theory Typerep
6imports String
7begin
8
9datatype typerep = Typerep String.literal "typerep list"
10
11class typerep =
12  fixes typerep :: "'a itself \<Rightarrow> typerep"
13begin
14
15definition typerep_of :: "'a \<Rightarrow> typerep" where
16  [simp]: "typerep_of x = typerep TYPE('a)"
17
18end
19
20syntax
21  "_TYPEREP" :: "type => logic"  ("(1TYPEREP/(1'(_')))")
22
23parse_translation \<open>
24  let
25    fun typerep_tr (*"_TYPEREP"*) [ty] =
26          Syntax.const @{const_syntax typerep} $
27            (Syntax.const @{syntax_const "_constrain"} $ Syntax.const @{const_syntax Pure.type} $
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
31\<close>
32
33typed_print_translation \<open>
34  let
35    fun typerep_tr' ctxt (*"typerep"*)
36            (Type (@{type_name fun}, [Type (@{type_name itself}, [T]), _]))
37            (Const (@{const_syntax Pure.type}, _) :: ts) =
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
42\<close>
43
44setup \<open>
45let
46
47fun add_typerep tyco thy =
48  let
49    val sorts = replicate (Sign.arity_number thy tyco) @{sort typerep};
50    val vs = Name.invent_names Name.context "'a" sorts;
51    val ty = Type (tyco, map TFree vs);
52    val lhs = Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep})
53      $ Free ("T", Term.itselfT ty);
54    val rhs = @{term Typerep} $ HOLogic.mk_literal tyco
55      $ HOLogic.mk_list @{typ typerep} (map (HOLogic.mk_typerep o TFree) vs);
56    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
57  in
58    thy
59    |> Class.instantiation ([tyco], vs, @{sort typerep})
60    |> `(fn lthy => Syntax.check_term lthy eq)
61    |-> (fn eq => Specification.definition NONE [] [] (Binding.empty_atts, eq))
62    |> snd
63    |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])
64  end;
65
66fun ensure_typerep tyco thy =
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}
69  then add_typerep tyco thy else thy;
70
71in
72
73add_typerep @{type_name fun}
74#> Typedef.interpretation (Local_Theory.background_theory o ensure_typerep)
75#> Code.type_interpretation ensure_typerep
76
77end
78\<close>
79
80lemma [code]:
81  "HOL.equal (Typerep tyco1 tys1) (Typerep tyco2 tys2) \<longleftrightarrow> HOL.equal tyco1 tyco2
82     \<and> list_all2 HOL.equal tys1 tys2"
83  by (auto simp add: eq_equal [symmetric] list_all2_eq [symmetric])
84
85lemma [code nbe]:
86  "HOL.equal (x :: typerep) x \<longleftrightarrow> True"
87  by (fact equal_refl)
88
89code_printing
90  type_constructor typerep \<rightharpoonup> (Eval) "Term.typ"
91| constant Typerep \<rightharpoonup> (Eval) "Term.Type/ (_, _)"
92
93code_reserved Eval Term
94
95hide_const (open) typerep Typerep
96
97end
98