author | wenzelm |
Sat, 27 Nov 2010 14:32:08 +0100 | |
changeset 40742 | dc6439c0b8b1 |
parent 38536 | 7e57a0dcbd4f |
child 41732 | 996b0c14a430 |
permissions | -rw-r--r-- |
11608 | 1 |
(* Title: HOL/Typedef.thy |
2 |
Author: Markus Wenzel, TU Munich |
|
11743 | 3 |
*) |
11608 | 4 |
|
11979 | 5 |
header {* HOL type definitions *} |
11608 | 6 |
|
15131 | 7 |
theory Typedef |
15140 | 8 |
imports Set |
38536
7e57a0dcbd4f
dropped SML typedef_codegen: does not fit to code equations for record operations any longer
haftmann
parents:
38393
diff
changeset
|
9 |
uses ("Tools/typedef.ML") |
15131 | 10 |
begin |
11608 | 11 |
|
23247 | 12 |
ML {* |
37863
7f113caabcf4
discontinued pervasive val theory = Thy_Info.get_theory -- prefer antiquotations in most situations;
wenzelm
parents:
31723
diff
changeset
|
13 |
structure HOL = struct val thy = @{theory HOL} end; |
23247 | 14 |
*} -- "belongs to theory HOL" |
15 |
||
13412 | 16 |
locale type_definition = |
17 |
fixes Rep and Abs and A |
|
18 |
assumes Rep: "Rep x \<in> A" |
|
19 |
and Rep_inverse: "Abs (Rep x) = x" |
|
20 |
and Abs_inverse: "y \<in> A ==> Rep (Abs y) = y" |
|
21 |
-- {* This will be axiomatized for each typedef! *} |
|
23247 | 22 |
begin |
11608 | 23 |
|
23247 | 24 |
lemma Rep_inject: |
13412 | 25 |
"(Rep x = Rep y) = (x = y)" |
26 |
proof |
|
27 |
assume "Rep x = Rep y" |
|
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
28 |
then have "Abs (Rep x) = Abs (Rep y)" by (simp only:) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
29 |
moreover have "Abs (Rep x) = x" by (rule Rep_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
30 |
moreover have "Abs (Rep y) = y" by (rule Rep_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
31 |
ultimately show "x = y" by simp |
13412 | 32 |
next |
33 |
assume "x = y" |
|
34 |
thus "Rep x = Rep y" by (simp only:) |
|
35 |
qed |
|
11608 | 36 |
|
23247 | 37 |
lemma Abs_inject: |
13412 | 38 |
assumes x: "x \<in> A" and y: "y \<in> A" |
39 |
shows "(Abs x = Abs y) = (x = y)" |
|
40 |
proof |
|
41 |
assume "Abs x = Abs y" |
|
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
42 |
then have "Rep (Abs x) = Rep (Abs y)" by (simp only:) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
43 |
moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
44 |
moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
45 |
ultimately show "x = y" by simp |
13412 | 46 |
next |
47 |
assume "x = y" |
|
48 |
thus "Abs x = Abs y" by (simp only:) |
|
11608 | 49 |
qed |
50 |
||
23247 | 51 |
lemma Rep_cases [cases set]: |
13412 | 52 |
assumes y: "y \<in> A" |
53 |
and hyp: "!!x. y = Rep x ==> P" |
|
54 |
shows P |
|
55 |
proof (rule hyp) |
|
56 |
from y have "Rep (Abs y) = y" by (rule Abs_inverse) |
|
57 |
thus "y = Rep (Abs y)" .. |
|
11608 | 58 |
qed |
59 |
||
23247 | 60 |
lemma Abs_cases [cases type]: |
13412 | 61 |
assumes r: "!!y. x = Abs y ==> y \<in> A ==> P" |
62 |
shows P |
|
63 |
proof (rule r) |
|
64 |
have "Abs (Rep x) = x" by (rule Rep_inverse) |
|
65 |
thus "x = Abs (Rep x)" .. |
|
66 |
show "Rep x \<in> A" by (rule Rep) |
|
11608 | 67 |
qed |
68 |
||
23247 | 69 |
lemma Rep_induct [induct set]: |
13412 | 70 |
assumes y: "y \<in> A" |
71 |
and hyp: "!!x. P (Rep x)" |
|
72 |
shows "P y" |
|
11608 | 73 |
proof - |
13412 | 74 |
have "P (Rep (Abs y))" by (rule hyp) |
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
75 |
moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
76 |
ultimately show "P y" by simp |
11608 | 77 |
qed |
78 |
||
23247 | 79 |
lemma Abs_induct [induct type]: |
13412 | 80 |
assumes r: "!!y. y \<in> A ==> P (Abs y)" |
81 |
shows "P x" |
|
11608 | 82 |
proof - |
13412 | 83 |
have "Rep x \<in> A" by (rule Rep) |
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
84 |
then have "P (Abs (Rep x))" by (rule r) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
85 |
moreover have "Abs (Rep x) = x" by (rule Rep_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
86 |
ultimately show "P x" by simp |
11608 | 87 |
qed |
88 |
||
27295 | 89 |
lemma Rep_range: "range Rep = A" |
24269 | 90 |
proof |
91 |
show "range Rep <= A" using Rep by (auto simp add: image_def) |
|
92 |
show "A <= range Rep" |
|
23433 | 93 |
proof |
94 |
fix x assume "x : A" |
|
24269 | 95 |
hence "x = Rep (Abs x)" by (rule Abs_inverse [symmetric]) |
96 |
thus "x : range Rep" by (rule range_eqI) |
|
23433 | 97 |
qed |
98 |
qed |
|
99 |
||
27295 | 100 |
lemma Abs_image: "Abs ` A = UNIV" |
101 |
proof |
|
102 |
show "Abs ` A <= UNIV" by (rule subset_UNIV) |
|
103 |
next |
|
104 |
show "UNIV <= Abs ` A" |
|
105 |
proof |
|
106 |
fix x |
|
107 |
have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric]) |
|
108 |
moreover have "Rep x : A" by (rule Rep) |
|
109 |
ultimately show "x : Abs ` A" by (rule image_eqI) |
|
110 |
qed |
|
111 |
qed |
|
112 |
||
23247 | 113 |
end |
114 |
||
31723
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
29797
diff
changeset
|
115 |
use "Tools/typedef.ML" setup Typedef.setup |
11608 | 116 |
|
117 |
end |