src/HOL/Library/RType.thy
author haftmann
Fri, 10 Oct 2008 06:45:53 +0200
changeset 28562 4e74209f113e
parent 28394 b9c8e3a12a98
permissions -rw-r--r--
`code func` now just `code`
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Library/RType.thy
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     2
    ID:         $Id$
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     4
*)
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     5
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     6
header {* Reflecting Pure types into HOL *}
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     7
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     8
theory RType
28228
7ebe8dc06cbb evaluation using code generator
haftmann
parents: 28084
diff changeset
     9
imports Plain "~~/src/HOL/List" "~~/src/HOL/Library/Code_Message"
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    10
begin
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    11
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    12
datatype typerep = Typerep message_string "typerep list"
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    13
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    14
class typerep =
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    15
  fixes typerep :: "'a\<Colon>{} itself \<Rightarrow> typerep"
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    16
begin
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    17
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    18
definition
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    19
  typerep_of :: "'a \<Rightarrow> typerep"
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    20
where
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    21
  [simp]: "typerep_of x = typerep TYPE('a)"
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    22
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    23
end
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    24
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    25
setup {*
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    26
let
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    27
  fun typerep_tr (*"_TYPEREP"*) [ty] =
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    28
        Lexicon.const @{const_syntax typerep} $ (Lexicon.const "_constrain" $ Lexicon.const "TYPE" $
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    29
          (Lexicon.const "itself" $ ty))
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    30
    | typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts);
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    31
  fun typerep_tr' show_sorts (*"typerep"*)
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    32
          (Type ("fun", [Type ("itself", [T]), _])) (Const (@{const_syntax TYPE}, _) :: ts) =
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    33
        Term.list_comb (Lexicon.const "_TYPEREP" $ Syntax.term_of_typ show_sorts T, ts)
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    34
    | typerep_tr' _ T ts = raise Match;
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    35
in
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    36
  Sign.add_syntax_i
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    37
    [("_TYPEREP", SimpleSyntax.read_typ "type => logic", Delimfix "(1TYPEREP/(1'(_')))")]
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    38
  #> Sign.add_trfuns ([], [("_TYPEREP", typerep_tr)], [], [])
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    39
  #> Sign.add_trfunsT [(@{const_syntax typerep}, typerep_tr')]
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    40
end
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    41
*}
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    42
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    43
ML {*
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    44
structure Typerep =
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    45
struct
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    46
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    47
fun mk f (Type (tyco, tys)) =
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    48
      @{term Typerep} $ Message_String.mk tyco
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    49
        $ HOLogic.mk_list @{typ typerep} (map (mk f) tys)
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    50
  | mk f (TFree v) =
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    51
      f v;
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    52
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    53
fun typerep ty =
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    54
  Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep})
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    55
    $ Logic.mk_type ty;
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    56
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    57
fun add_def tyco thy =
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    58
  let
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    59
    val sorts = replicate (Sign.arity_number thy tyco) @{sort typerep};
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    60
    val vs = Name.names Name.context "'a" sorts;
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    61
    val ty = Type (tyco, map TFree vs);
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    62
    val lhs = Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep})
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    63
      $ Free ("T", Term.itselfT ty);
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    64
    val rhs = mk (typerep o TFree) ty;
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    65
    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    66
  in
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    67
    thy
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    68
    |> TheoryTarget.instantiation ([tyco], vs, @{sort typerep})
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    69
    |> `(fn lthy => Syntax.check_term lthy eq)
28084
a05ca48ef263 type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents: 28083
diff changeset
    70
    |-> (fn eq => Specification.definition (NONE, (Attrib.no_binding, eq)))
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    71
    |> snd
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    72
    |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
28394
b9c8e3a12a98 LocalTheory.exit_global;
wenzelm
parents: 28346
diff changeset
    73
    |> LocalTheory.exit_global
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    74
  end;
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    75
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    76
fun perhaps_add_def tyco thy =
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    77
  let
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    78
    val inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep}
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    79
  in if inst then thy else add_def tyco thy end;
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    80
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    81
end;
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    82
*}
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    83
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    84
setup {*
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    85
  Typerep.add_def @{type_name prop}
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    86
  #> Typerep.add_def @{type_name fun}
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    87
  #> Typerep.add_def @{type_name itself}
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    88
  #> Typerep.add_def @{type_name bool}
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    89
  #> TypedefPackage.interpretation Typerep.perhaps_add_def
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    90
*}
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    91
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28394
diff changeset
    92
lemma [code]:
28346
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28335
diff changeset
    93
  "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
    94
     \<and> list_all2 eq_class.eq tys1 tys2"
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28335
diff changeset
    95
  by (auto simp add: equals_eq [symmetric] list_all2_eq [symmetric])
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    96
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    97
code_type typerep
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    98
  (SML "Term.typ")
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    99
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
   100
code_const Typerep
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
   101
  (SML "Term.Type/ (_, _)")
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
   102
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
   103
code_reserved SML Term
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
   104
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
   105
hide (open) const typerep Typerep
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
   106
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
   107
end