src/HOL/Typerep.thy
author haftmann
Thu, 29 Oct 2009 11:41:36 +0100
changeset 33318 ddd97d9dfbfb
parent 31723 f5cafe803b55
child 33384 1b5ba4e6a953
permissions -rw-r--r--
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31048
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31031
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     2
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     3
header {* Reflecting Pure types into HOL *}
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     4
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 28562
diff changeset
     5
theory Typerep
31048
ac146fc38b51 refined HOL string theories and corresponding ML fragments
haftmann
parents: 31031
diff changeset
     6
imports Plain String
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     7
begin
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
     8
31205
98370b26c2ce String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents: 31137
diff changeset
     9
datatype typerep = Typerep String.literal "typerep list"
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    10
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    11
class typerep =
31031
cbec39ebf8f2 class typerep inherits from type
haftmann
parents: 29574
diff changeset
    12
  fixes typerep :: "'a itself \<Rightarrow> typerep"
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    13
begin
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    14
29574
haftmann
parents: 28965
diff changeset
    15
definition typerep_of :: "'a \<Rightarrow> typerep" where
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    16
  [simp]: "typerep_of x = typerep TYPE('a)"
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    17
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    18
end
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    19
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    20
setup {*
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    21
let
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    22
  fun typerep_tr (*"_TYPEREP"*) [ty] =
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    23
        Lexicon.const @{const_syntax typerep} $ (Lexicon.const "_constrain" $ Lexicon.const "TYPE" $
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    24
          (Lexicon.const "itself" $ ty))
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    25
    | typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts);
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    26
  fun typerep_tr' show_sorts (*"typerep"*)
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    27
          (Type ("fun", [Type ("itself", [T]), _])) (Const (@{const_syntax TYPE}, _) :: ts) =
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    28
        Term.list_comb (Lexicon.const "_TYPEREP" $ Syntax.term_of_typ show_sorts T, ts)
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    29
    | typerep_tr' _ T ts = raise Match;
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    30
in
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    31
  Sign.add_syntax_i
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    32
    [("_TYPEREP", SimpleSyntax.read_typ "type => logic", Delimfix "(1TYPEREP/(1'(_')))")]
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    33
  #> Sign.add_trfuns ([], [("_TYPEREP", typerep_tr)], [], [])
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    34
  #> Sign.add_trfunsT [(@{const_syntax typerep}, typerep_tr')]
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    35
end
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    36
*}
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    37
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    38
setup {*
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    39
let
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    40
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    41
fun add_typerep tyco thy =
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    42
  let
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    43
    val sorts = replicate (Sign.arity_number thy tyco) @{sort typerep};
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    44
    val vs = Name.names Name.context "'a" sorts;
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    45
    val ty = Type (tyco, map TFree vs);
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    46
    val lhs = Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep})
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    47
      $ Free ("T", Term.itselfT ty);
31205
98370b26c2ce String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents: 31137
diff changeset
    48
    val rhs = @{term Typerep} $ HOLogic.mk_literal tyco
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    49
      $ HOLogic.mk_list @{typ typerep} (map (HOLogic.mk_typerep o TFree) vs);
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    50
    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    51
  in
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    52
    thy
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    53
    |> TheoryTarget.instantiation ([tyco], vs, @{sort typerep})
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    54
    |> `(fn lthy => Syntax.check_term lthy eq)
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28952
diff changeset
    55
    |-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    56
    |> snd
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    57
    |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    58
  end;
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    59
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    60
fun ensure_typerep tyco thy = if not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep})
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    61
  andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort type}
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    62
  then add_typerep tyco thy else thy;
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    63
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    64
in
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    65
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    66
add_typerep @{type_name fun}
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31205
diff changeset
    67
#> Typedef.interpretation ensure_typerep
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    68
#> Code.type_interpretation (ensure_typerep o fst)
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    69
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    70
end
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    71
*}
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    72
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28394
diff changeset
    73
lemma [code]:
28346
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28335
diff changeset
    74
  "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
    75
     \<and> list_all2 eq_class.eq tys1 tys2"
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28335
diff changeset
    76
  by (auto simp add: equals_eq [symmetric] list_all2_eq [symmetric])
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    77
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    78
code_type typerep
31031
cbec39ebf8f2 class typerep inherits from type
haftmann
parents: 29574
diff changeset
    79
  (Eval "Term.typ")
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    80
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    81
code_const Typerep
31031
cbec39ebf8f2 class typerep inherits from type
haftmann
parents: 29574
diff changeset
    82
  (Eval "Term.Type/ (_, _)")
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    83
31031
cbec39ebf8f2 class typerep inherits from type
haftmann
parents: 29574
diff changeset
    84
code_reserved Eval Term
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    85
28335
25326092cf9a renamed rtype to typerep
haftmann
parents: 28228
diff changeset
    86
hide (open) const typerep Typerep
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    87
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    88
end