| author | chaieb | 
| Mon, 09 Feb 2009 11:07:17 +0000 | |
| changeset 29835 | 62da280e5d0b | 
| parent 29574 | 5897b2688ccc | 
| child 31031 | cbec39ebf8f2 | 
| permissions | -rw-r--r-- | 
| 29574 | 1  | 
(* Title: HOL/Typerep.thy  | 
| 26168 | 2  | 
Author: Florian Haftmann, TU Muenchen  | 
3  | 
*)  | 
|
4  | 
||
5  | 
header {* Reflecting Pure types into HOL *}
 | 
|
6  | 
||
| 
28952
 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 
haftmann 
parents: 
28562 
diff
changeset
 | 
7  | 
theory Typerep  | 
| 
 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 
haftmann 
parents: 
28562 
diff
changeset
 | 
8  | 
imports Plain List Code_Message  | 
| 26168 | 9  | 
begin  | 
10  | 
||
| 28335 | 11  | 
datatype typerep = Typerep message_string "typerep list"  | 
| 26168 | 12  | 
|
| 28335 | 13  | 
class typerep =  | 
14  | 
  fixes typerep :: "'a\<Colon>{} itself \<Rightarrow> typerep"
 | 
|
| 26168 | 15  | 
begin  | 
16  | 
||
| 29574 | 17  | 
definition typerep_of :: "'a \<Rightarrow> typerep" where  | 
| 28335 | 18  | 
  [simp]: "typerep_of x = typerep TYPE('a)"
 | 
| 26168 | 19  | 
|
20  | 
end  | 
|
21  | 
||
22  | 
setup {*
 | 
|
23  | 
let  | 
|
| 28335 | 24  | 
fun typerep_tr (*"_TYPEREP"*) [ty] =  | 
25  | 
        Lexicon.const @{const_syntax typerep} $ (Lexicon.const "_constrain" $ Lexicon.const "TYPE" $
 | 
|
| 26168 | 26  | 
(Lexicon.const "itself" $ ty))  | 
| 28335 | 27  | 
    | typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts);
 | 
28  | 
fun typerep_tr' show_sorts (*"typerep"*)  | 
|
| 26168 | 29  | 
          (Type ("fun", [Type ("itself", [T]), _])) (Const (@{const_syntax TYPE}, _) :: ts) =
 | 
| 28335 | 30  | 
Term.list_comb (Lexicon.const "_TYPEREP" $ Syntax.term_of_typ show_sorts T, ts)  | 
31  | 
| typerep_tr' _ T ts = raise Match;  | 
|
| 26168 | 32  | 
in  | 
33  | 
Sign.add_syntax_i  | 
|
| 28335 | 34  | 
    [("_TYPEREP", SimpleSyntax.read_typ "type => logic", Delimfix "(1TYPEREP/(1'(_')))")]
 | 
35  | 
  #> Sign.add_trfuns ([], [("_TYPEREP", typerep_tr)], [], [])
 | 
|
36  | 
  #> Sign.add_trfunsT [(@{const_syntax typerep}, typerep_tr')]
 | 
|
| 26168 | 37  | 
end  | 
38  | 
*}  | 
|
39  | 
||
40  | 
ML {*
 | 
|
| 28335 | 41  | 
structure Typerep =  | 
| 26168 | 42  | 
struct  | 
43  | 
||
44  | 
fun mk f (Type (tyco, tys)) =  | 
|
| 28335 | 45  | 
      @{term Typerep} $ Message_String.mk tyco
 | 
46  | 
        $ HOLogic.mk_list @{typ typerep} (map (mk f) tys)
 | 
|
| 26168 | 47  | 
| mk f (TFree v) =  | 
48  | 
f v;  | 
|
49  | 
||
| 28335 | 50  | 
fun typerep ty =  | 
51  | 
  Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep})
 | 
|
| 26168 | 52  | 
$ Logic.mk_type ty;  | 
53  | 
||
54  | 
fun add_def tyco thy =  | 
|
55  | 
let  | 
|
| 28335 | 56  | 
    val sorts = replicate (Sign.arity_number thy tyco) @{sort typerep};
 | 
| 26168 | 57  | 
val vs = Name.names Name.context "'a" sorts;  | 
58  | 
val ty = Type (tyco, map TFree vs);  | 
|
| 28335 | 59  | 
    val lhs = Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep})
 | 
| 26168 | 60  | 
      $ Free ("T", Term.itselfT ty);
 | 
| 28335 | 61  | 
val rhs = mk (typerep o TFree) ty;  | 
| 26168 | 62  | 
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));  | 
63  | 
in  | 
|
64  | 
thy  | 
|
| 28335 | 65  | 
    |> TheoryTarget.instantiation ([tyco], vs, @{sort typerep})
 | 
| 26168 | 66  | 
|> `(fn lthy => Syntax.check_term lthy eq)  | 
| 28965 | 67  | 
|-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))  | 
| 26168 | 68  | 
|> snd  | 
69  | 
|> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))  | 
|
| 28394 | 70  | 
|> LocalTheory.exit_global  | 
| 26168 | 71  | 
end;  | 
72  | 
||
73  | 
fun perhaps_add_def tyco thy =  | 
|
74  | 
let  | 
|
| 28335 | 75  | 
    val inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep}
 | 
| 26168 | 76  | 
in if inst then thy else add_def tyco thy end;  | 
77  | 
||
78  | 
end;  | 
|
79  | 
*}  | 
|
80  | 
||
81  | 
setup {*
 | 
|
| 28335 | 82  | 
  Typerep.add_def @{type_name prop}
 | 
83  | 
  #> Typerep.add_def @{type_name fun}
 | 
|
84  | 
  #> Typerep.add_def @{type_name itself}
 | 
|
85  | 
  #> Typerep.add_def @{type_name bool}
 | 
|
86  | 
#> TypedefPackage.interpretation Typerep.perhaps_add_def  | 
|
| 26168 | 87  | 
*}  | 
88  | 
||
| 28562 | 89  | 
lemma [code]:  | 
| 
28346
 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 
haftmann 
parents: 
28335 
diff
changeset
 | 
90  | 
"eq_class.eq (Typerep tyco1 tys1) (Typerep tyco2 tys2) \<longleftrightarrow> eq_class.eq tyco1 tyco2  | 
| 
 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 
haftmann 
parents: 
28335 
diff
changeset
 | 
91  | 
\<and> list_all2 eq_class.eq tys1 tys2"  | 
| 
 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 
haftmann 
parents: 
28335 
diff
changeset
 | 
92  | 
by (auto simp add: equals_eq [symmetric] list_all2_eq [symmetric])  | 
| 26168 | 93  | 
|
| 28335 | 94  | 
code_type typerep  | 
| 26168 | 95  | 
(SML "Term.typ")  | 
96  | 
||
| 28335 | 97  | 
code_const Typerep  | 
| 26168 | 98  | 
(SML "Term.Type/ (_, _)")  | 
99  | 
||
100  | 
code_reserved SML Term  | 
|
101  | 
||
| 28335 | 102  | 
hide (open) const typerep Typerep  | 
| 26168 | 103  | 
|
104  | 
end  |