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