Up to index of Isabelle/HOL
theory Typedef(* Title: HOL/Typedef.thy ID: $Id: Typedef.thy,v 1.24 2008/05/07 08:56:50 berghofe Exp $ Author: Markus Wenzel, TU Munich *) header {* HOL type definitions *} theory Typedef imports Set uses ("Tools/typedef_package.ML") ("Tools/typecopy_package.ML") ("Tools/typedef_codegen.ML") begin ML {* structure HOL = struct val thy = theory "HOL" end; *} -- "belongs to theory HOL" locale type_definition = fixes Rep and Abs and A assumes Rep: "Rep x ∈ A" and Rep_inverse: "Abs (Rep x) = x" and Abs_inverse: "y ∈ A ==> Rep (Abs y) = y" -- {* This will be axiomatized for each typedef! *} begin lemma Rep_inject: "(Rep x = Rep y) = (x = y)" proof assume "Rep x = Rep y" then have "Abs (Rep x) = Abs (Rep y)" by (simp only:) moreover have "Abs (Rep x) = x" by (rule Rep_inverse) moreover have "Abs (Rep y) = y" by (rule Rep_inverse) ultimately show "x = y" by simp next assume "x = y" thus "Rep x = Rep y" by (simp only:) qed lemma Abs_inject: assumes x: "x ∈ A" and y: "y ∈ A" shows "(Abs x = Abs y) = (x = y)" proof assume "Abs x = Abs y" then have "Rep (Abs x) = Rep (Abs y)" by (simp only:) moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse) moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse) ultimately show "x = y" by simp next assume "x = y" thus "Abs x = Abs y" by (simp only:) qed lemma Rep_cases [cases set]: assumes y: "y ∈ A" and hyp: "!!x. y = Rep x ==> P" shows P proof (rule hyp) from y have "Rep (Abs y) = y" by (rule Abs_inverse) thus "y = Rep (Abs y)" .. qed lemma Abs_cases [cases type]: assumes r: "!!y. x = Abs y ==> y ∈ A ==> P" shows P proof (rule r) have "Abs (Rep x) = x" by (rule Rep_inverse) thus "x = Abs (Rep x)" .. show "Rep x ∈ A" by (rule Rep) qed lemma Rep_induct [induct set]: assumes y: "y ∈ A" and hyp: "!!x. P (Rep x)" shows "P y" proof - have "P (Rep (Abs y))" by (rule hyp) moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse) ultimately show "P y" by simp qed lemma Abs_induct [induct type]: assumes r: "!!y. y ∈ A ==> P (Abs y)" shows "P x" proof - have "Rep x ∈ A" by (rule Rep) then have "P (Abs (Rep x))" by (rule r) moreover have "Abs (Rep x) = x" by (rule Rep_inverse) ultimately show "P x" by simp qed lemma Rep_range: shows "range Rep = A" proof show "range Rep <= A" using Rep by (auto simp add: image_def) show "A <= range Rep" proof fix x assume "x : A" hence "x = Rep (Abs x)" by (rule Abs_inverse [symmetric]) thus "x : range Rep" by (rule range_eqI) qed qed end use "Tools/typedef_package.ML" use "Tools/typecopy_package.ML" use "Tools/typedef_codegen.ML" setup {* TypedefPackage.setup #> TypecopyPackage.setup #> TypedefCodegen.setup *} text {* This class is just a workaround for classes without parameters; it shall disappear as soon as possible. *} class itself = type + fixes itself :: "'a itself" setup {* let fun add_itself tyco thy = let val vs = Name.names Name.context "'a" (replicate (Sign.arity_number thy tyco) @{sort type}); val ty = Type (tyco, map TFree vs); val lhs = Const (@{const_name itself}, Term.itselfT ty); val rhs = Logic.mk_type ty; val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); in thy |> TheoryTarget.instantiation ([tyco], vs, @{sort itself}) |> `(fn lthy => Syntax.check_term lthy eq) |-> (fn eq => Specification.definition (NONE, (("", []), eq))) |> snd |> Class.prove_instantiation_instance (K (Class.intro_classes_tac [])) |> LocalTheory.exit |> ProofContext.theory_of end in TypedefPackage.interpretation add_itself end *} instantiation bool :: itself begin definition "itself = TYPE(bool)" instance .. end instantiation "fun" :: ("type", "type") itself begin definition "itself = TYPE('a => 'b)" instance .. end hide (open) const itself end
lemma Rep_inject:
(Rep x = Rep y) = (x = y)
lemma Abs_inject:
[| x ∈ A; y ∈ A |] ==> (Abs x = Abs y) = (x = y)
lemma Rep_cases:
[| y ∈ A; !!x. y = Rep x ==> P |] ==> P
lemma Abs_cases:
(!!y. [| x = Abs y; y ∈ A |] ==> P) ==> P
lemma Rep_induct:
[| y ∈ A; !!x. P (Rep x) |] ==> P y
lemma Abs_induct:
(!!y. y ∈ A ==> P (Abs y)) ==> P x
lemma Rep_range:
range Rep = A