author | wenzelm |
Sat, 10 Jan 2009 21:32:30 +0100 | |
changeset 29435 | a5f84ac14609 |
parent 29056 | dc08e3990c77 |
child 29608 | 564ea783ace8 |
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 |
20426 | 9 |
uses |
10 |
("Tools/typedef_package.ML") |
|
11 |
("Tools/typecopy_package.ML") |
|
12 |
("Tools/typedef_codegen.ML") |
|
15131 | 13 |
begin |
11608 | 14 |
|
23247 | 15 |
ML {* |
16 |
structure HOL = struct val thy = theory "HOL" end; |
|
17 |
*} -- "belongs to theory HOL" |
|
18 |
||
13412 | 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! *} |
|
23247 | 25 |
begin |
11608 | 26 |
|
23247 | 27 |
lemma Rep_inject: |
13412 | 28 |
"(Rep x = Rep y) = (x = y)" |
29 |
proof |
|
30 |
assume "Rep x = Rep y" |
|
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
31 |
then have "Abs (Rep x) = Abs (Rep y)" by (simp only:) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
32 |
moreover have "Abs (Rep x) = x" by (rule Rep_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
33 |
moreover have "Abs (Rep y) = y" by (rule Rep_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
34 |
ultimately show "x = y" by simp |
13412 | 35 |
next |
36 |
assume "x = y" |
|
37 |
thus "Rep x = Rep y" by (simp only:) |
|
38 |
qed |
|
11608 | 39 |
|
23247 | 40 |
lemma Abs_inject: |
13412 | 41 |
assumes x: "x \<in> A" and y: "y \<in> A" |
42 |
shows "(Abs x = Abs y) = (x = y)" |
|
43 |
proof |
|
44 |
assume "Abs x = Abs y" |
|
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
45 |
then have "Rep (Abs x) = Rep (Abs y)" by (simp only:) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
46 |
moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
47 |
moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
48 |
ultimately show "x = y" by simp |
13412 | 49 |
next |
50 |
assume "x = y" |
|
51 |
thus "Abs x = Abs y" by (simp only:) |
|
11608 | 52 |
qed |
53 |
||
23247 | 54 |
lemma Rep_cases [cases set]: |
13412 | 55 |
assumes y: "y \<in> A" |
56 |
and hyp: "!!x. y = Rep x ==> P" |
|
57 |
shows P |
|
58 |
proof (rule hyp) |
|
59 |
from y have "Rep (Abs y) = y" by (rule Abs_inverse) |
|
60 |
thus "y = Rep (Abs y)" .. |
|
11608 | 61 |
qed |
62 |
||
23247 | 63 |
lemma Abs_cases [cases type]: |
13412 | 64 |
assumes r: "!!y. x = Abs y ==> y \<in> A ==> P" |
65 |
shows P |
|
66 |
proof (rule r) |
|
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) |
|
11608 | 70 |
qed |
71 |
||
23247 | 72 |
lemma Rep_induct [induct set]: |
13412 | 73 |
assumes y: "y \<in> A" |
74 |
and hyp: "!!x. P (Rep x)" |
|
75 |
shows "P y" |
|
11608 | 76 |
proof - |
13412 | 77 |
have "P (Rep (Abs y))" by (rule hyp) |
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
78 |
moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
79 |
ultimately show "P y" by simp |
11608 | 80 |
qed |
81 |
||
23247 | 82 |
lemma Abs_induct [induct type]: |
13412 | 83 |
assumes r: "!!y. y \<in> A ==> P (Abs y)" |
84 |
shows "P x" |
|
11608 | 85 |
proof - |
13412 | 86 |
have "Rep x \<in> A" by (rule Rep) |
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
87 |
then have "P (Abs (Rep x))" by (rule r) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
88 |
moreover have "Abs (Rep x) = x" by (rule Rep_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
89 |
ultimately show "P x" by simp |
11608 | 90 |
qed |
91 |
||
27295 | 92 |
lemma Rep_range: "range Rep = A" |
24269 | 93 |
proof |
94 |
show "range Rep <= A" using Rep by (auto simp add: image_def) |
|
95 |
show "A <= range Rep" |
|
23433 | 96 |
proof |
97 |
fix x assume "x : A" |
|
24269 | 98 |
hence "x = Rep (Abs x)" by (rule Abs_inverse [symmetric]) |
99 |
thus "x : range Rep" by (rule range_eqI) |
|
23433 | 100 |
qed |
101 |
qed |
|
102 |
||
27295 | 103 |
lemma Abs_image: "Abs ` A = UNIV" |
104 |
proof |
|
105 |
show "Abs ` A <= UNIV" by (rule subset_UNIV) |
|
106 |
next |
|
107 |
show "UNIV <= Abs ` A" |
|
108 |
proof |
|
109 |
fix x |
|
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) |
|
113 |
qed |
|
114 |
qed |
|
115 |
||
23247 | 116 |
end |
117 |
||
29056 | 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 |
|
11608 | 121 |
|
15260 | 122 |
|
26151
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
123 |
text {* This class is just a workaround for classes without parameters; |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
124 |
it shall disappear as soon as possible. *} |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
125 |
|
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
126 |
class itself = type + |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
127 |
fixes itself :: "'a itself" |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
128 |
|
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
129 |
setup {* |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
130 |
let fun add_itself tyco thy = |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
131 |
let |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
132 |
val vs = Name.names Name.context "'a" |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
133 |
(replicate (Sign.arity_number thy tyco) @{sort type}); |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
134 |
val ty = Type (tyco, map TFree vs); |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
135 |
val lhs = Const (@{const_name itself}, Term.itselfT ty); |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
136 |
val rhs = Logic.mk_type ty; |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
137 |
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
138 |
in |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
139 |
thy |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
140 |
|> TheoryTarget.instantiation ([tyco], vs, @{sort itself}) |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
141 |
|> `(fn lthy => Syntax.check_term lthy eq) |
28965 | 142 |
|-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq))) |
26151
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
143 |
|> snd |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
144 |
|> Class.prove_instantiation_instance (K (Class.intro_classes_tac [])) |
28394 | 145 |
|> LocalTheory.exit_global |
26151
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
146 |
end |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
147 |
in TypedefPackage.interpretation add_itself end |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
148 |
*} |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
149 |
|
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
150 |
instantiation bool :: itself |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
151 |
begin |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
152 |
|
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
153 |
definition "itself = TYPE(bool)" |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
154 |
|
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
155 |
instance .. |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
156 |
|
11608 | 157 |
end |
26151
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
158 |
|
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
159 |
instantiation "fun" :: ("type", "type") itself |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
160 |
begin |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
161 |
|
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
162 |
definition "itself = TYPE('a \<Rightarrow> 'b)" |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
163 |
|
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
164 |
instance .. |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
165 |
|
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
166 |
end |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
167 |
|
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
168 |
hide (open) const itself |
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
169 |
|
4a9b8f15ce7f
class itself works around a problem with class interpretation in class finite
haftmann
parents:
25535
diff
changeset
|
170 |
end |