add generic phantom type
authorAndreas Lochbihler
Thu, 28 Jun 2012 09:14:57 +0200
changeset 48163 f0ecc1550998
parent 48162 5717466d4633
child 48164 e97369f20c30
add generic phantom type
src/HOL/IsaMakefile
src/HOL/Library/Phantom_Type.thy
--- a/src/HOL/IsaMakefile	Wed Jun 27 17:52:07 2012 +0200
+++ b/src/HOL/IsaMakefile	Thu Jun 28 09:14:57 2012 +0200
@@ -469,17 +469,17 @@
   Library/Numeral_Type.thy Library/Old_Recdef.thy			\
   Library/OptionalSugar.thy Library/Order_Relation.thy			\
   Library/Permutation.thy Library/Permutations.thy			\
-  Library/Poly_Deriv.thy Library/Polynomial.thy				\
-  Library/Predicate_Compile_Quickcheck.thy Library/Preorder.thy		\
-  Library/Product_Vector.thy Library/Product_ord.thy			\
-  Library/Product_plus.thy Library/Product_Lattice.thy			\
-  Library/Quickcheck_Types.thy Library/Quotient_List.thy		\
-  Library/Quotient_Option.thy Library/Quotient_Product.thy		\
-  Library/Quotient_Set.thy Library/Quotient_Sum.thy			\
-  Library/Quotient_Syntax.thy Library/Quotient_Type.thy Library/RBT.thy	\
-  Library/RBT_Impl.thy Library/RBT_Mapping.thy Library/README.html	\
-  Library/Saturated.thy Library/Set_Algebras.thy			\
-  Library/State_Monad.thy Library/Ramsey.thy				\
+  Library/Phantom_Type.thy Library/Poly_Deriv.thy			\
+  Library/Polynomial.thy Library/Predicate_Compile_Quickcheck.thy	\
+  Library/Preorder.thy Library/Product_Vector.thy			\
+  Library/Product_ord.thy Library/Product_plus.thy 			\
+  Library/Product_Lattice.thy Library/Quickcheck_Types.thy 		\
+  Library/Quotient_List.thy Library/Quotient_Option.thy 		\
+  Library/Quotient_Product.thy Library/Quotient_Set.thy 		\
+  Library/Quotient_Sum.thy Library/Quotient_Syntax.thy 			\
+  Library/Quotient_Type.thy Library/RBT.thy Library/RBT_Impl.thy	\
+  Library/RBT_Mapping.thy Library/README.html Library/Saturated.thy	\
+  Library/Set_Algebras.thy Library/State_Monad.thy Library/Ramsey.thy	\
   Library/Reflection.thy Library/Sublist_Order.thy			\
   Library/Sum_of_Squares.thy Library/Sum_of_Squares/sos_wrapper.ML	\
   Library/Sum_of_Squares/sum_of_squares.ML Library/Target_Numeral.thy	\
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Phantom_Type.thy	Thu Jun 28 09:14:57 2012 +0200
@@ -0,0 +1,39 @@
+(*  Title:      HOL/Library/Phantom_Type.thy
+    Author:     Andreas Lochbihler
+*)
+
+header {* A generic phantom type *}
+
+theory Phantom_Type imports "~~/src/HOL/Main" begin
+
+datatype ('a, 'b) phantom = phantom 'b
+
+primrec of_phantom :: "('a, 'b) phantom \<Rightarrow> 'b" 
+where "of_phantom (phantom x) = x"
+
+lemma of_phantom_phantom [simp]: "phantom (of_phantom x) = x"
+by(cases x) simp
+
+lemma type_definition_phantom': "type_definition of_phantom phantom UNIV"
+by(unfold_locales) simp_all
+
+setup_lifting (no_code) type_definition_phantom'
+
+lemma phantom_comp_of_phantom [simp]: "phantom \<circ> of_phantom = id"
+  and of_phantom_comp_phantom [simp]: "of_phantom \<circ> phantom = id"
+by(simp_all add: o_def id_def)
+
+syntax "_Phantom" :: "type \<Rightarrow> logic" ("(1Phantom/(1'(_')))")
+translations
+  "Phantom('t)" => "CONST phantom :: _ \<Rightarrow> ('t, _) phantom"
+
+typed_print_translation (advanced) {*
+let
+  fun phantom_tr' ctxt 
+      (Type (@{type_name fun}, [_, Type (@{type_name phantom}, [T, _])])) ts =
+    Term.list_comb (Syntax.const @{syntax_const "_Phantom"} $ Syntax_Phases.term_of_typ ctxt T, ts)
+  | phantom_tr' _ _ _ = raise Match;
+in [(@{const_syntax phantom}, phantom_tr')] end
+*}
+
+end