author | paulson |
Wed, 25 Jul 2001 17:58:26 +0200 | |
changeset 11454 | 7514e5e21cb8 |
parent 11083 | d8fda557e476 |
permissions | -rw-r--r-- |
1475 | 1 |
(* Title: HOL/subset.thy |
923 | 2 |
ID: $Id$ |
1475 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
923 | 4 |
Copyright 1994 University of Cambridge |
10276 | 5 |
|
6 |
Subset lemmas and HOL type definitions. |
|
923 | 7 |
*) |
8 |
||
7705 | 9 |
theory subset = Set |
10276 | 10 |
files "Tools/induct_attrib.ML" ("Tools/typedef_package.ML"): |
7705 | 11 |
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
12 |
(** belongs to theory Ord **) |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
13 |
|
9895 | 14 |
theorems linorder_cases [case_names less equal greater] = |
15 |
linorder_less_split |
|
16 |
||
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
17 |
(* Courtesy of Stephan Merz *) |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
18 |
lemma Least_mono: |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
19 |
"[| mono (f::'a::order => 'b::order); EX x:S. ALL y:S. x <= y |] |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
20 |
==> (LEAST y. y : f`S) = f(LEAST x. x : S)" |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
21 |
apply clarify |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
22 |
apply (erule_tac P = "%x. x : S" in LeastI2) |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
23 |
apply fast |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
24 |
apply (rule LeastI2) |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
25 |
apply (auto elim: monoD intro!: order_antisym) |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
26 |
done |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
27 |
|
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
28 |
|
9895 | 29 |
(*belongs to theory Set*) |
30 |
setup Rulify.setup |
|
7717
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents:
7705
diff
changeset
|
31 |
|
10276 | 32 |
|
33 |
section {* HOL type definitions *} |
|
34 |
||
35 |
constdefs |
|
36 |
type_definition :: "('a => 'b) => ('b => 'a) => 'b set => bool" |
|
37 |
"type_definition Rep Abs A == |
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
38 |
(\\<forall>x. Rep x \\<in> A) \\<and> |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
39 |
(\\<forall>x. Abs (Rep x) = x) \\<and> |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
40 |
(\\<forall>y \\<in> A. Rep (Abs y) = y)" |
10276 | 41 |
-- {* This will be stated as an axiom for each typedef! *} |
42 |
||
10290 | 43 |
lemma type_definitionI [intro]: |
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
44 |
"(!!x. Rep x \\<in> A) ==> |
10291 | 45 |
(!!x. Abs (Rep x) = x) ==> |
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
46 |
(!!y. y \\<in> A ==> Rep (Abs y) = y) ==> |
10290 | 47 |
type_definition Rep Abs A" |
48 |
by (unfold type_definition_def) blast |
|
49 |
||
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
50 |
theorem Rep: "type_definition Rep Abs A ==> Rep x \\<in> A" |
10276 | 51 |
by (unfold type_definition_def) blast |
52 |
||
53 |
theorem Rep_inverse: "type_definition Rep Abs A ==> Abs (Rep x) = x" |
|
54 |
by (unfold type_definition_def) blast |
|
55 |
||
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
56 |
theorem Abs_inverse: "type_definition Rep Abs A ==> y \\<in> A ==> Rep (Abs y) = y" |
10276 | 57 |
by (unfold type_definition_def) blast |
58 |
||
59 |
theorem Rep_inject: "type_definition Rep Abs A ==> (Rep x = Rep y) = (x = y)" |
|
60 |
proof - |
|
61 |
assume tydef: "type_definition Rep Abs A" |
|
62 |
show ?thesis |
|
63 |
proof |
|
64 |
assume "Rep x = Rep y" |
|
65 |
hence "Abs (Rep x) = Abs (Rep y)" by (simp only:) |
|
66 |
thus "x = y" by (simp only: Rep_inverse [OF tydef]) |
|
67 |
next |
|
68 |
assume "x = y" |
|
69 |
thus "Rep x = Rep y" by simp |
|
70 |
qed |
|
71 |
qed |
|
72 |
||
10284 | 73 |
theorem Abs_inject: |
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
74 |
"type_definition Rep Abs A ==> x \\<in> A ==> y \\<in> A ==> (Abs x = Abs y) = (x = y)" |
10284 | 75 |
proof - |
76 |
assume tydef: "type_definition Rep Abs A" |
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
77 |
assume x: "x \\<in> A" and y: "y \\<in> A" |
10284 | 78 |
show ?thesis |
79 |
proof |
|
80 |
assume "Abs x = Abs y" |
|
81 |
hence "Rep (Abs x) = Rep (Abs y)" by simp |
|
11083 | 82 |
moreover from x have "Rep (Abs x) = x" by (rule Abs_inverse [OF tydef]) |
83 |
moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef]) |
|
10284 | 84 |
ultimately show "x = y" by (simp only:) |
85 |
next |
|
86 |
assume "x = y" |
|
87 |
thus "Abs x = Abs y" by simp |
|
88 |
qed |
|
89 |
qed |
|
90 |
||
10276 | 91 |
theorem Rep_cases: |
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
92 |
"type_definition Rep Abs A ==> y \\<in> A ==> (!!x. y = Rep x ==> P) ==> P" |
10276 | 93 |
proof - |
94 |
assume tydef: "type_definition Rep Abs A" |
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
95 |
assume y: "y \\<in> A" and r: "(!!x. y = Rep x ==> P)" |
10276 | 96 |
show P |
97 |
proof (rule r) |
|
98 |
from y have "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef]) |
|
99 |
thus "y = Rep (Abs y)" .. |
|
100 |
qed |
|
101 |
qed |
|
102 |
||
103 |
theorem Abs_cases: |
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
104 |
"type_definition Rep Abs A ==> (!!y. x = Abs y ==> y \\<in> A ==> P) ==> P" |
10276 | 105 |
proof - |
106 |
assume tydef: "type_definition Rep Abs A" |
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
107 |
assume r: "!!y. x = Abs y ==> y \\<in> A ==> P" |
10276 | 108 |
show P |
109 |
proof (rule r) |
|
110 |
have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef]) |
|
111 |
thus "x = Abs (Rep x)" .. |
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
112 |
show "Rep x \\<in> A" by (rule Rep [OF tydef]) |
10276 | 113 |
qed |
114 |
qed |
|
115 |
||
116 |
theorem Rep_induct: |
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
117 |
"type_definition Rep Abs A ==> y \\<in> A ==> (!!x. P (Rep x)) ==> P y" |
10276 | 118 |
proof - |
119 |
assume tydef: "type_definition Rep Abs A" |
|
120 |
assume "!!x. P (Rep x)" hence "P (Rep (Abs y))" . |
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
121 |
moreover assume "y \\<in> A" hence "Rep (Abs y) = y" by (rule Abs_inverse [OF tydef]) |
10276 | 122 |
ultimately show "P y" by (simp only:) |
123 |
qed |
|
124 |
||
125 |
theorem Abs_induct: |
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
126 |
"type_definition Rep Abs A ==> (!!y. y \\<in> A ==> P (Abs y)) ==> P x" |
10276 | 127 |
proof - |
128 |
assume tydef: "type_definition Rep Abs A" |
|
11454
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
129 |
assume r: "!!y. y \\<in> A ==> P (Abs y)" |
7514e5e21cb8
Hilbert restructuring: Wellfounded_Relations no longer needs Hilbert_Choice
paulson
parents:
11083
diff
changeset
|
130 |
have "Rep x \\<in> A" by (rule Rep [OF tydef]) |
10276 | 131 |
hence "P (Abs (Rep x))" by (rule r) |
132 |
moreover have "Abs (Rep x) = x" by (rule Rep_inverse [OF tydef]) |
|
133 |
ultimately show "P x" by (simp only:) |
|
134 |
qed |
|
135 |
||
136 |
setup InductAttrib.setup |
|
137 |
use "Tools/typedef_package.ML" |
|
138 |
||
7705 | 139 |
end |