src/HOL/Library/Phantom_Type.thy
author wenzelm
Sun, 13 Dec 2020 23:11:41 +0100
changeset 72907 3883f536d84d
parent 69593 3dda49e08b9d
permissions -rw-r--r--
unused (see 29566b6810f7);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48163
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
     1
(*  Title:      HOL/Library/Phantom_Type.thy
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
     2
    Author:     Andreas Lochbihler
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
     3
*)
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
     4
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
     5
section \<open>A generic phantom type\<close>
48163
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
     6
51542
738598beeb26 tuned imports;
wenzelm
parents: 51378
diff changeset
     7
theory Phantom_Type
738598beeb26 tuned imports;
wenzelm
parents: 51378
diff changeset
     8
imports Main
738598beeb26 tuned imports;
wenzelm
parents: 51378
diff changeset
     9
begin
48163
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    10
58378
cf6f16bc11a7 use selector
blanchet
parents: 58310
diff changeset
    11
datatype ('a, 'b) phantom = phantom (of_phantom: 'b)
48163
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    12
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    13
lemma type_definition_phantom': "type_definition of_phantom phantom UNIV"
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    14
by(unfold_locales) simp_all
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    15
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    16
lemma phantom_comp_of_phantom [simp]: "phantom \<circ> of_phantom = id"
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    17
  and of_phantom_comp_phantom [simp]: "of_phantom \<circ> phantom = id"
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    18
by(simp_all add: o_def id_def)
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    19
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    20
syntax "_Phantom" :: "type \<Rightarrow> logic" ("(1Phantom/(1'(_')))")
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    21
translations
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    22
  "Phantom('t)" => "CONST phantom :: _ \<Rightarrow> ('t, _) phantom"
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    23
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    24
typed_print_translation \<open>
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51542
diff changeset
    25
  let
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 60500
diff changeset
    26
    fun phantom_tr' ctxt (Type (\<^type_name>\<open>fun\<close>, [_, Type (\<^type_name>\<open>phantom\<close>, [T, _])])) ts =
52147
wenzelm
parents: 52143
diff changeset
    27
          list_comb
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 60500
diff changeset
    28
            (Syntax.const \<^syntax_const>\<open>_Phantom\<close> $ Syntax_Phases.term_of_typ ctxt T, ts)
52147
wenzelm
parents: 52143
diff changeset
    29
      | phantom_tr' _ _ _ = raise Match;
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 60500
diff changeset
    30
  in [(\<^const_syntax>\<open>phantom\<close>, phantom_tr')] end
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    31
\<close>
48163
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    32
58383
09a2c3e08ec2 add lemma
Andreas Lochbihler
parents: 58310
diff changeset
    33
lemma of_phantom_inject [simp]:
09a2c3e08ec2 add lemma
Andreas Lochbihler
parents: 58310
diff changeset
    34
  "of_phantom x = of_phantom y \<longleftrightarrow> x = y"
09a2c3e08ec2 add lemma
Andreas Lochbihler
parents: 58310
diff changeset
    35
by(cases x y rule: phantom.exhaust[case_product phantom.exhaust]) simp
09a2c3e08ec2 add lemma
Andreas Lochbihler
parents: 58310
diff changeset
    36
48163
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    37
end