src/HOL/Library/Phantom_Type.thy
author hoelzl
Tue, 12 Nov 2013 19:28:50 +0100
changeset 54408 67dec4ccaabd
parent 52147 9943f8067f11
child 58249 180f1b3508ed
permissions -rw-r--r--
equation when indicator function equals 0 or 1
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
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
     5
header {* A generic phantom type *}
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
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    11
datatype ('a, 'b) phantom = phantom 'b
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    12
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    13
primrec of_phantom :: "('a, 'b) phantom \<Rightarrow> 'b" 
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    14
where "of_phantom (phantom x) = x"
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    15
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    16
lemma of_phantom_phantom [simp]: "phantom (of_phantom x) = x"
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    17
by(cases x) simp
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    18
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    19
lemma type_definition_phantom': "type_definition of_phantom phantom UNIV"
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    20
by(unfold_locales) simp_all
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    21
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    22
lemma phantom_comp_of_phantom [simp]: "phantom \<circ> of_phantom = id"
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    23
  and of_phantom_comp_phantom [simp]: "of_phantom \<circ> phantom = id"
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    24
by(simp_all add: o_def id_def)
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    25
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    26
syntax "_Phantom" :: "type \<Rightarrow> logic" ("(1Phantom/(1'(_')))")
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    27
translations
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    28
  "Phantom('t)" => "CONST phantom :: _ \<Rightarrow> ('t, _) phantom"
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    29
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51542
diff changeset
    30
typed_print_translation {*
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51542
diff changeset
    31
  let
52147
wenzelm
parents: 52143
diff changeset
    32
    fun phantom_tr' ctxt (Type (@{type_name fun}, [_, Type (@{type_name phantom}, [T, _])])) ts =
wenzelm
parents: 52143
diff changeset
    33
          list_comb
wenzelm
parents: 52143
diff changeset
    34
            (Syntax.const @{syntax_const "_Phantom"} $ Syntax_Phases.term_of_typ ctxt T, ts)
wenzelm
parents: 52143
diff changeset
    35
      | phantom_tr' _ _ _ = raise Match;
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51542
diff changeset
    36
  in [(@{const_syntax phantom}, phantom_tr')] end
48163
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    37
*}
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    38
f0ecc1550998 add generic phantom type
Andreas Lochbihler
parents:
diff changeset
    39
end