src/HOL/Typerep.thy
author wenzelm
Tue, 05 Nov 2019 14:16:16 +0100
changeset 71046 b8aeeedf7e68
parent 69593 3dda49e08b9d
child 74345 e5ff77db6f38
permissions -rw-r--r--
support for Linux user management;
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
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59498
diff changeset
     3
section \<open>Reflecting Pure types into HOL\<close>
26168
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
51112
da97167e03f7 abandoned theory Plain
haftmann
parents: 48272
diff changeset
     6
imports 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
58310
91ea607a34d8 updated news
blanchet
parents: 58152
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
35115
446c5063e4fd modernized translations;
wenzelm
parents: 33553
diff changeset
    20
syntax
446c5063e4fd modernized translations;
wenzelm
parents: 33553
diff changeset
    21
  "_TYPEREP" :: "type => logic"  ("(1TYPEREP/(1'(_')))")
446c5063e4fd modernized translations;
wenzelm
parents: 33553
diff changeset
    22
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59498
diff changeset
    23
parse_translation \<open>
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51112
diff changeset
    24
  let
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51112
diff changeset
    25
    fun typerep_tr (*"_TYPEREP"*) [ty] =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    26
          Syntax.const \<^const_syntax>\<open>typerep\<close> $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    27
            (Syntax.const \<^syntax_const>\<open>_constrain\<close> $ Syntax.const \<^const_syntax>\<open>Pure.type\<close> $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    28
              (Syntax.const \<^type_syntax>\<open>itself\<close> $ ty))
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51112
diff changeset
    29
      | typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts);
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    30
  in [(\<^syntax_const>\<open>_TYPEREP\<close>, K typerep_tr)] end
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59498
diff changeset
    31
\<close>
35115
446c5063e4fd modernized translations;
wenzelm
parents: 33553
diff changeset
    32
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59498
diff changeset
    33
typed_print_translation \<open>
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51112
diff changeset
    34
  let
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51112
diff changeset
    35
    fun typerep_tr' ctxt (*"typerep"*)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    36
            (Type (\<^type_name>\<open>fun\<close>, [Type (\<^type_name>\<open>itself\<close>, [T]), _]))
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    37
            (Const (\<^const_syntax>\<open>Pure.type\<close>, _) :: ts) =
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51112
diff changeset
    38
          Term.list_comb
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    39
            (Syntax.const \<^syntax_const>\<open>_TYPEREP\<close> $ Syntax_Phases.term_of_typ ctxt T, ts)
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51112
diff changeset
    40
      | typerep_tr' _ T ts = raise Match;
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    41
  in [(\<^const_syntax>\<open>typerep\<close>, typerep_tr')] end
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59498
diff changeset
    42
\<close>
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    43
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59498
diff changeset
    44
setup \<open>
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    45
let
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    46
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    47
fun add_typerep tyco thy =
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    48
  let
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    49
    val sorts = replicate (Sign.arity_number thy tyco) \<^sort>\<open>typerep\<close>;
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 42247
diff changeset
    50
    val vs = Name.invent_names Name.context "'a" sorts;
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    51
    val ty = Type (tyco, map TFree vs);
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    52
    val lhs = Const (\<^const_name>\<open>typerep\<close>, Term.itselfT ty --> \<^typ>\<open>typerep\<close>)
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    53
      $ Free ("T", Term.itselfT ty);
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    54
    val rhs = \<^term>\<open>Typerep\<close> $ HOLogic.mk_literal tyco
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    55
      $ HOLogic.mk_list \<^typ>\<open>typerep\<close> (map (HOLogic.mk_typerep o TFree) vs);
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    56
    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    57
  in
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    58
    thy
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    59
    |> Class.instantiation ([tyco], vs, \<^sort>\<open>typerep\<close>)
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    60
    |> `(fn lthy => Syntax.check_term lthy eq)
63352
4eaf35781b23 tuned signature;
wenzelm
parents: 63180
diff changeset
    61
    |-> (fn eq => Specification.definition NONE [] [] (Binding.empty_atts, eq))
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    62
    |> snd
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
    63
    |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt [])
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    64
  end;
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    65
45693
bbd2c7ffc02c tuned layout;
wenzelm
parents: 43329
diff changeset
    66
fun ensure_typerep tyco thy =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    67
  if not (Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\<open>typerep\<close>)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    68
    andalso Sorts.has_instance (Sign.classes_of thy) tyco \<^sort>\<open>type\<close>
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    69
  then add_typerep tyco thy else thy;
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    70
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    71
in
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    72
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 66330
diff changeset
    73
add_typerep \<^type_name>\<open>fun\<close>
58662
5963cdbad926 module Interpretation is superseded by Plugin;
wenzelm
parents: 58310
diff changeset
    74
#> Typedef.interpretation (Local_Theory.background_theory o ensure_typerep)
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 63352
diff changeset
    75
#> Code.type_interpretation ensure_typerep
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    76
31137
cd3aafc1c631 tuned construction of typerep instances
haftmann
parents: 31048
diff changeset
    77
end
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59498
diff changeset
    78
\<close>
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    79
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28394
diff changeset
    80
lemma [code]:
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38348
diff changeset
    81
  "HOL.equal (Typerep tyco1 tys1) (Typerep tyco2 tys2) \<longleftrightarrow> HOL.equal tyco1 tyco2
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38348
diff changeset
    82
     \<and> list_all2 HOL.equal tys1 tys2"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38348
diff changeset
    83
  by (auto simp add: eq_equal [symmetric] list_all2_eq [symmetric])
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38348
diff changeset
    84
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38348
diff changeset
    85
lemma [code nbe]:
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38348
diff changeset
    86
  "HOL.equal (x :: typerep) x \<longleftrightarrow> True"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38348
diff changeset
    87
  by (fact equal_refl)
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    88
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
    89
code_printing
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
    90
  type_constructor typerep \<rightharpoonup> (Eval) "Term.typ"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
    91
| constant Typerep \<rightharpoonup> (Eval) "Term.Type/ (_, _)"
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    92
31031
cbec39ebf8f2 class typerep inherits from type
haftmann
parents: 29574
diff changeset
    93
code_reserved Eval Term
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    94
36176
3fe7e97ccca8 replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents: 35430
diff changeset
    95
hide_const (open) typerep Typerep
26168
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    96
3bd9ac4e0b97 added theory for reflected types
haftmann
parents:
diff changeset
    97
end