author | wenzelm |
Mon, 11 Apr 2016 15:26:58 +0200 | |
changeset 62954 | c5d0fdc260fa |
parent 61799 | 4cf66f21b764 |
child 63434 | c956d995bec6 |
permissions | -rw-r--r-- |
11608 | 1 |
(* Title: HOL/Typedef.thy |
2 |
Author: Markus Wenzel, TU Munich |
|
11743 | 3 |
*) |
11608 | 4 |
|
60758 | 5 |
section \<open>HOL type definitions\<close> |
11608 | 6 |
|
15131 | 7 |
theory Typedef |
15140 | 8 |
imports Set |
46950
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
46947
diff
changeset
|
9 |
keywords "typedef" :: thy_goal and "morphisms" |
15131 | 10 |
begin |
11608 | 11 |
|
13412 | 12 |
locale type_definition = |
13 |
fixes Rep and Abs and A |
|
14 |
assumes Rep: "Rep x \<in> A" |
|
15 |
and Rep_inverse: "Abs (Rep x) = x" |
|
61102 | 16 |
and Abs_inverse: "y \<in> A \<Longrightarrow> Rep (Abs y) = y" |
61799 | 17 |
\<comment> \<open>This will be axiomatized for each typedef!\<close> |
23247 | 18 |
begin |
11608 | 19 |
|
61102 | 20 |
lemma Rep_inject: "Rep x = Rep y \<longleftrightarrow> x = y" |
13412 | 21 |
proof |
22 |
assume "Rep x = Rep y" |
|
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
23 |
then have "Abs (Rep x) = Abs (Rep y)" by (simp only:) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
24 |
moreover have "Abs (Rep x) = x" by (rule Rep_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
25 |
moreover have "Abs (Rep y) = y" by (rule Rep_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
26 |
ultimately show "x = y" by simp |
13412 | 27 |
next |
28 |
assume "x = y" |
|
61102 | 29 |
then show "Rep x = Rep y" by (simp only:) |
13412 | 30 |
qed |
11608 | 31 |
|
23247 | 32 |
lemma Abs_inject: |
61102 | 33 |
assumes "x \<in> A" and "y \<in> A" |
34 |
shows "Abs x = Abs y \<longleftrightarrow> x = y" |
|
13412 | 35 |
proof |
36 |
assume "Abs x = Abs y" |
|
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
37 |
then have "Rep (Abs x) = Rep (Abs y)" by (simp only:) |
61102 | 38 |
moreover from \<open>x \<in> A\<close> have "Rep (Abs x) = x" by (rule Abs_inverse) |
39 |
moreover from \<open>y \<in> A\<close> have "Rep (Abs y) = y" by (rule Abs_inverse) |
|
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
40 |
ultimately show "x = y" by simp |
13412 | 41 |
next |
42 |
assume "x = y" |
|
61102 | 43 |
then show "Abs x = Abs y" by (simp only:) |
11608 | 44 |
qed |
45 |
||
23247 | 46 |
lemma Rep_cases [cases set]: |
61102 | 47 |
assumes "y \<in> A" |
48 |
and hyp: "\<And>x. y = Rep x \<Longrightarrow> P" |
|
13412 | 49 |
shows P |
50 |
proof (rule hyp) |
|
61102 | 51 |
from \<open>y \<in> A\<close> have "Rep (Abs y) = y" by (rule Abs_inverse) |
52 |
then show "y = Rep (Abs y)" .. |
|
11608 | 53 |
qed |
54 |
||
23247 | 55 |
lemma Abs_cases [cases type]: |
61102 | 56 |
assumes r: "\<And>y. x = Abs y \<Longrightarrow> y \<in> A \<Longrightarrow> P" |
13412 | 57 |
shows P |
58 |
proof (rule r) |
|
59 |
have "Abs (Rep x) = x" by (rule Rep_inverse) |
|
61102 | 60 |
then show "x = Abs (Rep x)" .. |
13412 | 61 |
show "Rep x \<in> A" by (rule Rep) |
11608 | 62 |
qed |
63 |
||
23247 | 64 |
lemma Rep_induct [induct set]: |
13412 | 65 |
assumes y: "y \<in> A" |
61102 | 66 |
and hyp: "\<And>x. P (Rep x)" |
13412 | 67 |
shows "P y" |
11608 | 68 |
proof - |
13412 | 69 |
have "P (Rep (Abs y))" by (rule hyp) |
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
70 |
moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
71 |
ultimately show "P y" by simp |
11608 | 72 |
qed |
73 |
||
23247 | 74 |
lemma Abs_induct [induct type]: |
61102 | 75 |
assumes r: "\<And>y. y \<in> A \<Longrightarrow> P (Abs y)" |
13412 | 76 |
shows "P x" |
11608 | 77 |
proof - |
13412 | 78 |
have "Rep x \<in> A" by (rule Rep) |
23710
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
79 |
then have "P (Abs (Rep x))" by (rule r) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
80 |
moreover have "Abs (Rep x) = x" by (rule Rep_inverse) |
a8ac2305eaf2
removed proof dependency on transitivity theorems
haftmann
parents:
23433
diff
changeset
|
81 |
ultimately show "P x" by simp |
11608 | 82 |
qed |
83 |
||
27295 | 84 |
lemma Rep_range: "range Rep = A" |
24269 | 85 |
proof |
61102 | 86 |
show "range Rep \<subseteq> A" using Rep by (auto simp add: image_def) |
87 |
show "A \<subseteq> range Rep" |
|
23433 | 88 |
proof |
61102 | 89 |
fix x assume "x \<in> A" |
90 |
then have "x = Rep (Abs x)" by (rule Abs_inverse [symmetric]) |
|
91 |
then show "x \<in> range Rep" by (rule range_eqI) |
|
23433 | 92 |
qed |
93 |
qed |
|
94 |
||
27295 | 95 |
lemma Abs_image: "Abs ` A = UNIV" |
96 |
proof |
|
61102 | 97 |
show "Abs ` A \<subseteq> UNIV" by (rule subset_UNIV) |
98 |
show "UNIV \<subseteq> Abs ` A" |
|
27295 | 99 |
proof |
100 |
fix x |
|
101 |
have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric]) |
|
61102 | 102 |
moreover have "Rep x \<in> A" by (rule Rep) |
103 |
ultimately show "x \<in> Abs ` A" by (rule image_eqI) |
|
27295 | 104 |
qed |
105 |
qed |
|
106 |
||
23247 | 107 |
end |
108 |
||
58239
1c5bc387bd4c
added flag to 'typedef' to allow concealed definitions
blanchet
parents:
48891
diff
changeset
|
109 |
ML_file "Tools/typedef.ML" |
11608 | 110 |
|
111 |
end |