1 (* Title: HOL/Typedef.thy
2 Author: Markus Wenzel, TU Munich
5 header {* HOL type definitions *}
10 ("Tools/typedef_package.ML")
11 ("Tools/typecopy_package.ML")
12 ("Tools/typedef_codegen.ML")
16 structure HOL = struct val thy = theory "HOL" end;
17 *} -- "belongs to theory HOL"
19 locale type_definition =
20 fixes Rep and Abs and A
21 assumes Rep: "Rep x \<in> A"
22 and Rep_inverse: "Abs (Rep x) = x"
23 and Abs_inverse: "y \<in> A ==> Rep (Abs y) = y"
24 -- {* This will be axiomatized for each typedef! *}
28 "(Rep x = Rep y) = (x = y)"
30 assume "Rep x = Rep y"
31 then have "Abs (Rep x) = Abs (Rep y)" by (simp only:)
32 moreover have "Abs (Rep x) = x" by (rule Rep_inverse)
33 moreover have "Abs (Rep y) = y" by (rule Rep_inverse)
34 ultimately show "x = y" by simp
37 thus "Rep x = Rep y" by (simp only:)
41 assumes x: "x \<in> A" and y: "y \<in> A"
42 shows "(Abs x = Abs y) = (x = y)"
44 assume "Abs x = Abs y"
45 then have "Rep (Abs x) = Rep (Abs y)" by (simp only:)
46 moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse)
47 moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse)
48 ultimately show "x = y" by simp
51 thus "Abs x = Abs y" by (simp only:)
54 lemma Rep_cases [cases set]:
55 assumes y: "y \<in> A"
56 and hyp: "!!x. y = Rep x ==> P"
59 from y have "Rep (Abs y) = y" by (rule Abs_inverse)
60 thus "y = Rep (Abs y)" ..
63 lemma Abs_cases [cases type]:
64 assumes r: "!!y. x = Abs y ==> y \<in> A ==> P"
67 have "Abs (Rep x) = x" by (rule Rep_inverse)
68 thus "x = Abs (Rep x)" ..
69 show "Rep x \<in> A" by (rule Rep)
72 lemma Rep_induct [induct set]:
73 assumes y: "y \<in> A"
74 and hyp: "!!x. P (Rep x)"
77 have "P (Rep (Abs y))" by (rule hyp)
78 moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse)
79 ultimately show "P y" by simp
82 lemma Abs_induct [induct type]:
83 assumes r: "!!y. y \<in> A ==> P (Abs y)"
86 have "Rep x \<in> A" by (rule Rep)
87 then have "P (Abs (Rep x))" by (rule r)
88 moreover have "Abs (Rep x) = x" by (rule Rep_inverse)
89 ultimately show "P x" by simp
92 lemma Rep_range: "range Rep = A"
94 show "range Rep <= A" using Rep by (auto simp add: image_def)
98 hence "x = Rep (Abs x)" by (rule Abs_inverse [symmetric])
99 thus "x : range Rep" by (rule range_eqI)
103 lemma Abs_image: "Abs ` A = UNIV"
105 show "Abs ` A <= UNIV" by (rule subset_UNIV)
107 show "UNIV <= Abs ` A"
110 have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric])
111 moreover have "Rep x : A" by (rule Rep)
112 ultimately show "x : Abs ` A" by (rule image_eqI)
118 use "Tools/typedef_package.ML" setup TypedefPackage.setup
119 use "Tools/typecopy_package.ML" setup TypecopyPackage.setup
120 use "Tools/typedef_codegen.ML" setup TypedefCodegen.setup