Andreas@48163: (* Title: HOL/Library/Phantom_Type.thy Andreas@48163: Author: Andreas Lochbihler Andreas@48163: *) Andreas@48163: Andreas@48163: header {* A generic phantom type *} Andreas@48163: wenzelm@51542: theory Phantom_Type wenzelm@51542: imports Main wenzelm@51542: begin Andreas@48163: blanchet@58378: datatype ('a, 'b) phantom = phantom (of_phantom: 'b) Andreas@48163: Andreas@48163: lemma type_definition_phantom': "type_definition of_phantom phantom UNIV" Andreas@48163: by(unfold_locales) simp_all Andreas@48163: Andreas@48163: lemma phantom_comp_of_phantom [simp]: "phantom \ of_phantom = id" Andreas@48163: and of_phantom_comp_phantom [simp]: "of_phantom \ phantom = id" Andreas@48163: by(simp_all add: o_def id_def) Andreas@48163: Andreas@48163: syntax "_Phantom" :: "type \ logic" ("(1Phantom/(1'(_')))") Andreas@48163: translations Andreas@48163: "Phantom('t)" => "CONST phantom :: _ \ ('t, _) phantom" Andreas@48163: wenzelm@52143: typed_print_translation {* wenzelm@52143: let wenzelm@52147: fun phantom_tr' ctxt (Type (@{type_name fun}, [_, Type (@{type_name phantom}, [T, _])])) ts = wenzelm@52147: list_comb wenzelm@52147: (Syntax.const @{syntax_const "_Phantom"} $ Syntax_Phases.term_of_typ ctxt T, ts) wenzelm@52147: | phantom_tr' _ _ _ = raise Match; wenzelm@52143: in [(@{const_syntax phantom}, phantom_tr')] end Andreas@48163: *} Andreas@48163: Andreas@58383: lemma of_phantom_inject [simp]: Andreas@58383: "of_phantom x = of_phantom y \ x = y" Andreas@58383: by(cases x y rule: phantom.exhaust[case_product phantom.exhaust]) simp Andreas@58383: Andreas@48163: end