| author | wenzelm | 
| Mon, 06 Feb 2023 16:11:05 +0100 | |
| changeset 77215 | 6cc3b131f761 | 
| parent 75669 | 43f5dfb7fa35 | 
| child 82248 | e8c96013ea8a | 
| permissions | -rw-r--r-- | 
| 
54552
 
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
 
blanchet 
parents: 
54551 
diff
changeset
 | 
1  | 
(* Title: HOL/Order_Relation.thy  | 
| 
 
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
 
blanchet 
parents: 
54551 
diff
changeset
 | 
2  | 
Author: Tobias Nipkow  | 
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
3  | 
Author: Andrei Popescu, TU Muenchen  | 
| 
54552
 
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
 
blanchet 
parents: 
54551 
diff
changeset
 | 
4  | 
*)  | 
| 26273 | 5  | 
|
| 60758 | 6  | 
section \<open>Orders as Relations\<close>  | 
| 26273 | 7  | 
|
8  | 
theory Order_Relation  | 
|
| 55027 | 9  | 
imports Wfrec  | 
| 26273 | 10  | 
begin  | 
11  | 
||
| 63572 | 12  | 
subsection \<open>Orders on a set\<close>  | 
| 26295 | 13  | 
|
| 30198 | 14  | 
definition "preorder_on A r \<equiv> refl_on A r \<and> trans r"  | 
| 26295 | 15  | 
|
16  | 
definition "partial_order_on A r \<equiv> preorder_on A r \<and> antisym r"  | 
|
| 26273 | 17  | 
|
| 26295 | 18  | 
definition "linear_order_on A r \<equiv> partial_order_on A r \<and> total_on A r"  | 
19  | 
||
20  | 
definition "strict_linear_order_on A r \<equiv> trans r \<and> irrefl r \<and> total_on A r"  | 
|
21  | 
||
22  | 
definition "well_order_on A r \<equiv> linear_order_on A r \<and> wf(r - Id)"  | 
|
| 26273 | 23  | 
|
| 26295 | 24  | 
lemmas order_on_defs =  | 
25  | 
preorder_on_def partial_order_on_def linear_order_on_def  | 
|
26  | 
strict_linear_order_on_def well_order_on_def  | 
|
27  | 
||
| 
68745
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
28  | 
lemma partial_order_onD:  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
29  | 
assumes "partial_order_on A r" shows "refl_on A r" and "trans r" and "antisym r"  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
30  | 
using assms unfolding partial_order_on_def preorder_on_def by auto  | 
| 26273 | 31  | 
|
| 26295 | 32  | 
lemma preorder_on_empty[simp]: "preorder_on {} {}"
 | 
| 63572 | 33  | 
by (simp add: preorder_on_def trans_def)  | 
| 26295 | 34  | 
|
35  | 
lemma partial_order_on_empty[simp]: "partial_order_on {} {}"
 | 
|
| 63572 | 36  | 
by (simp add: partial_order_on_def)  | 
| 26273 | 37  | 
|
| 26295 | 38  | 
lemma lnear_order_on_empty[simp]: "linear_order_on {} {}"
 | 
| 63572 | 39  | 
by (simp add: linear_order_on_def)  | 
| 26295 | 40  | 
|
41  | 
lemma well_order_on_empty[simp]: "well_order_on {} {}"
 | 
|
| 63572 | 42  | 
by (simp add: well_order_on_def)  | 
| 26295 | 43  | 
|
| 26273 | 44  | 
|
| 63572 | 45  | 
lemma preorder_on_converse[simp]: "preorder_on A (r\<inverse>) = preorder_on A r"  | 
46  | 
by (simp add: preorder_on_def)  | 
|
| 26295 | 47  | 
|
| 63572 | 48  | 
lemma partial_order_on_converse[simp]: "partial_order_on A (r\<inverse>) = partial_order_on A r"  | 
49  | 
by (simp add: partial_order_on_def)  | 
|
| 26273 | 50  | 
|
| 63572 | 51  | 
lemma linear_order_on_converse[simp]: "linear_order_on A (r\<inverse>) = linear_order_on A r"  | 
52  | 
by (simp add: linear_order_on_def)  | 
|
| 26295 | 53  | 
|
| 26273 | 54  | 
|
| 70180 | 55  | 
lemma partial_order_on_acyclic:  | 
56  | 
"partial_order_on A r \<Longrightarrow> acyclic (r - Id)"  | 
|
57  | 
by (simp add: acyclic_irrefl partial_order_on_def preorder_on_def trans_diff_Id)  | 
|
58  | 
||
59  | 
lemma partial_order_on_well_order_on:  | 
|
60  | 
"finite r \<Longrightarrow> partial_order_on A r \<Longrightarrow> wf (r - Id)"  | 
|
61  | 
by (simp add: finite_acyclic_wf partial_order_on_acyclic)  | 
|
62  | 
||
| 63572 | 63  | 
lemma strict_linear_order_on_diff_Id: "linear_order_on A r \<Longrightarrow> strict_linear_order_on A (r - Id)"  | 
64  | 
by (simp add: order_on_defs trans_diff_Id)  | 
|
| 26295 | 65  | 
|
| 
63563
 
0bcd79da075b
prefer [simp] over [iff] as [iff] break HOL-UNITY
 
Andreas Lochbihler 
parents: 
63561 
diff
changeset
 | 
66  | 
lemma linear_order_on_singleton [simp]: "linear_order_on {x} {(x, x)}"
 | 
| 63572 | 67  | 
by (simp add: order_on_defs)  | 
| 
63561
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
68  | 
|
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
69  | 
lemma linear_order_on_acyclic:  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
70  | 
assumes "linear_order_on A r"  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
71  | 
shows "acyclic (r - Id)"  | 
| 63572 | 72  | 
using strict_linear_order_on_diff_Id[OF assms]  | 
73  | 
by (auto simp add: acyclic_irrefl strict_linear_order_on_def)  | 
|
| 
63561
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
74  | 
|
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
75  | 
lemma linear_order_on_well_order_on:  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
76  | 
assumes "finite r"  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
77  | 
shows "linear_order_on A r \<longleftrightarrow> well_order_on A r"  | 
| 63572 | 78  | 
unfolding well_order_on_def  | 
79  | 
using assms finite_acyclic_wf[OF _ linear_order_on_acyclic, of r] by blast  | 
|
| 
63561
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
80  | 
|
| 26295 | 81  | 
|
| 63572 | 82  | 
subsection \<open>Orders on the field\<close>  | 
| 26273 | 83  | 
|
| 30198 | 84  | 
abbreviation "Refl r \<equiv> refl_on (Field r) r"  | 
| 26295 | 85  | 
|
86  | 
abbreviation "Preorder r \<equiv> preorder_on (Field r) r"  | 
|
87  | 
||
88  | 
abbreviation "Partial_order r \<equiv> partial_order_on (Field r) r"  | 
|
| 26273 | 89  | 
|
| 26295 | 90  | 
abbreviation "Total r \<equiv> total_on (Field r) r"  | 
91  | 
||
92  | 
abbreviation "Linear_order r \<equiv> linear_order_on (Field r) r"  | 
|
93  | 
||
94  | 
abbreviation "Well_order r \<equiv> well_order_on (Field r) r"  | 
|
95  | 
||
| 26273 | 96  | 
|
97  | 
lemma subset_Image_Image_iff:  | 
|
| 63572 | 98  | 
"Preorder r \<Longrightarrow> A \<subseteq> Field r \<Longrightarrow> B \<subseteq> Field r \<Longrightarrow>  | 
99  | 
r `` A \<subseteq> r `` B \<longleftrightarrow> (\<forall>a\<in>A.\<exists>b\<in>B. (b, a) \<in> r)"  | 
|
100  | 
apply (simp add: preorder_on_def refl_on_def Image_def subset_eq)  | 
|
101  | 
apply (simp only: trans_def)  | 
|
102  | 
apply fast  | 
|
103  | 
done  | 
|
| 26273 | 104  | 
|
105  | 
lemma subset_Image1_Image1_iff:  | 
|
| 63572 | 106  | 
  "Preorder r \<Longrightarrow> a \<in> Field r \<Longrightarrow> b \<in> Field r \<Longrightarrow> r `` {a} \<subseteq> r `` {b} \<longleftrightarrow> (b, a) \<in> r"
 | 
107  | 
by (simp add: subset_Image_Image_iff)  | 
|
| 26273 | 108  | 
|
109  | 
lemma Refl_antisym_eq_Image1_Image1_iff:  | 
|
| 63572 | 110  | 
assumes "Refl r"  | 
111  | 
and as: "antisym r"  | 
|
112  | 
and abf: "a \<in> Field r" "b \<in> Field r"  | 
|
| 
54552
 
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
 
blanchet 
parents: 
54551 
diff
changeset
 | 
113  | 
  shows "r `` {a} = r `` {b} \<longleftrightarrow> a = b"
 | 
| 63572 | 114  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
| 
54552
 
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
 
blanchet 
parents: 
54551 
diff
changeset
 | 
115  | 
proof  | 
| 63572 | 116  | 
assume ?lhs  | 
117  | 
then have *: "\<And>x. (a, x) \<in> r \<longleftrightarrow> (b, x) \<in> r"  | 
|
118  | 
by (simp add: set_eq_iff)  | 
|
119  | 
have "(a, a) \<in> r" "(b, b) \<in> r" using \<open>Refl r\<close> abf by (simp_all add: refl_on_def)  | 
|
120  | 
then have "(a, b) \<in> r" "(b, a) \<in> r" using *[of a] *[of b] by simp_all  | 
|
121  | 
then show ?rhs  | 
|
122  | 
using \<open>antisym r\<close>[unfolded antisym_def] by blast  | 
|
123  | 
next  | 
|
124  | 
assume ?rhs  | 
|
125  | 
then show ?lhs by fast  | 
|
126  | 
qed  | 
|
| 26273 | 127  | 
|
128  | 
lemma Partial_order_eq_Image1_Image1_iff:  | 
|
| 63572 | 129  | 
  "Partial_order r \<Longrightarrow> a \<in> Field r \<Longrightarrow> b \<in> Field r \<Longrightarrow> r `` {a} = r `` {b} \<longleftrightarrow> a = b"
 | 
130  | 
by (auto simp: order_on_defs Refl_antisym_eq_Image1_Image1_iff)  | 
|
| 26295 | 131  | 
|
| 52182 | 132  | 
lemma Total_Id_Field:  | 
| 63572 | 133  | 
assumes "Total r"  | 
134  | 
and not_Id: "\<not> r \<subseteq> Id"  | 
|
135  | 
shows "Field r = Field (r - Id)"  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
136  | 
proof -  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
137  | 
have "Field r \<subseteq> Field (r - Id)"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
138  | 
proof (rule subsetI)  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
139  | 
fix a assume *: "a \<in> Field r"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
140  | 
    from not_Id have "r \<noteq> {}" by fast
 | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
141  | 
with not_Id obtain b and c where "b \<noteq> c \<and> (b,c) \<in> r" by auto  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
142  | 
    then have "b \<noteq> c \<and> {b, c} \<subseteq> Field r" by (auto simp: Field_def)
 | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
143  | 
with * obtain d where "d \<in> Field r" "d \<noteq> a" by auto  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
144  | 
with * \<open>Total r\<close> have "(a, d) \<in> r \<or> (d, a) \<in> r" by (simp add: total_on_def)  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
145  | 
with \<open>d \<noteq> a\<close> show "a \<in> Field (r - Id)" unfolding Field_def by blast  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
146  | 
qed  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
147  | 
then show ?thesis  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
148  | 
using mono_Field[of "r - Id" r] Diff_subset[of r Id] by auto  | 
| 52182 | 149  | 
qed  | 
150  | 
||
| 
68745
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
151  | 
subsection\<open>Relations given by a predicate and the field\<close>  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
152  | 
|
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
153  | 
definition relation_of :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set"
 | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
154  | 
  where "relation_of P A \<equiv> { (a, b) \<in> A \<times> A. P a b }"
 | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
155  | 
|
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
156  | 
lemma Field_relation_of:  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
157  | 
assumes "refl_on A (relation_of P A)" shows "Field (relation_of P A) = A"  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
158  | 
using assms unfolding refl_on_def Field_def by auto  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
159  | 
|
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
160  | 
lemma partial_order_on_relation_ofI:  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
161  | 
assumes refl: "\<And>a. a \<in> A \<Longrightarrow> P a a"  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
162  | 
and trans: "\<And>a b c. \<lbrakk> a \<in> A; b \<in> A; c \<in> A \<rbrakk> \<Longrightarrow> P a b \<Longrightarrow> P b c \<Longrightarrow> P a c"  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
163  | 
and antisym: "\<And>a b. \<lbrakk> a \<in> A; b \<in> A \<rbrakk> \<Longrightarrow> P a b \<Longrightarrow> P b a \<Longrightarrow> a = b"  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
164  | 
shows "partial_order_on A (relation_of P A)"  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
165  | 
proof -  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
166  | 
from refl have "refl_on A (relation_of P A)"  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
167  | 
unfolding refl_on_def relation_of_def by auto  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
168  | 
moreover have "trans (relation_of P A)" and "antisym (relation_of P A)"  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
169  | 
unfolding relation_of_def  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
170  | 
by (auto intro: transI dest: trans, auto intro: antisymI dest: antisym)  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
171  | 
ultimately show ?thesis  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
172  | 
unfolding partial_order_on_def preorder_on_def by simp  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
173  | 
qed  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
174  | 
|
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
175  | 
lemma Partial_order_relation_ofI:  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
176  | 
assumes "partial_order_on A (relation_of P A)" shows "Partial_order (relation_of P A)"  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
177  | 
using Field_relation_of assms partial_order_on_def preorder_on_def by fastforce  | 
| 
 
345ce5f262ea
Zorn's lemma for relations defined by predicates
 
paulson <lp15@cam.ac.uk> 
parents: 
68484 
diff
changeset
 | 
178  | 
|
| 26295 | 179  | 
|
| 63572 | 180  | 
subsection \<open>Orders on a type\<close>  | 
| 26295 | 181  | 
|
182  | 
abbreviation "strict_linear_order \<equiv> strict_linear_order_on UNIV"  | 
|
183  | 
||
184  | 
abbreviation "linear_order \<equiv> linear_order_on UNIV"  | 
|
185  | 
||
| 54551 | 186  | 
abbreviation "well_order \<equiv> well_order_on UNIV"  | 
| 26273 | 187  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
188  | 
|
| 60758 | 189  | 
subsection \<open>Order-like relations\<close>  | 
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
190  | 
|
| 63572 | 191  | 
text \<open>  | 
192  | 
In this subsection, we develop basic concepts and results pertaining  | 
|
193  | 
to order-like relations, i.e., to reflexive and/or transitive and/or symmetric and/or  | 
|
194  | 
total relations. We also further define upper and lower bounds operators.  | 
|
195  | 
\<close>  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
196  | 
|
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
197  | 
|
| 60758 | 198  | 
subsubsection \<open>Auxiliaries\<close>  | 
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
199  | 
|
| 63572 | 200  | 
lemma refl_on_domain: "refl_on A r \<Longrightarrow> (a, b) \<in> r \<Longrightarrow> a \<in> A \<and> b \<in> A"  | 
201  | 
by (auto simp add: refl_on_def)  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
202  | 
|
| 63572 | 203  | 
corollary well_order_on_domain: "well_order_on A r \<Longrightarrow> (a, b) \<in> r \<Longrightarrow> a \<in> A \<and> b \<in> A"  | 
204  | 
by (auto simp add: refl_on_domain order_on_defs)  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
205  | 
|
| 63572 | 206  | 
lemma well_order_on_Field: "well_order_on A r \<Longrightarrow> A = Field r"  | 
207  | 
by (auto simp add: refl_on_def Field_def order_on_defs)  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
208  | 
|
| 63572 | 209  | 
lemma well_order_on_Well_order: "well_order_on A r \<Longrightarrow> A = Field r \<and> Well_order r"  | 
210  | 
using well_order_on_Field [of A] by auto  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
211  | 
|
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
212  | 
lemma Total_subset_Id:  | 
| 63572 | 213  | 
assumes "Total r"  | 
214  | 
and "r \<subseteq> Id"  | 
|
215  | 
  shows "r = {} \<or> (\<exists>a. r = {(a, a)})"
 | 
|
216  | 
proof -  | 
|
217  | 
  have "\<exists>a. r = {(a, a)}" if "r \<noteq> {}"
 | 
|
218  | 
proof -  | 
|
219  | 
from that obtain a b where ab: "(a, b) \<in> r" by fast  | 
|
220  | 
with \<open>r \<subseteq> Id\<close> have "a = b" by blast  | 
|
221  | 
with ab have aa: "(a, a) \<in> r" by simp  | 
|
222  | 
have "a = c \<and> a = d" if "(c, d) \<in> r" for c d  | 
|
223  | 
proof -  | 
|
224  | 
      from that have "{a, c, d} \<subseteq> Field r"
 | 
|
225  | 
using ab unfolding Field_def by blast  | 
|
226  | 
then have "((a, c) \<in> r \<or> (c, a) \<in> r \<or> a = c) \<and> ((a, d) \<in> r \<or> (d, a) \<in> r \<or> a = d)"  | 
|
227  | 
using \<open>Total r\<close> unfolding total_on_def by blast  | 
|
228  | 
with \<open>r \<subseteq> Id\<close> show ?thesis by blast  | 
|
229  | 
qed  | 
|
230  | 
    then have "r \<subseteq> {(a, a)}" by auto
 | 
|
231  | 
with aa show ?thesis by blast  | 
|
232  | 
qed  | 
|
233  | 
then show ?thesis by blast  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
234  | 
qed  | 
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
235  | 
|
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
236  | 
lemma Linear_order_in_diff_Id:  | 
| 63572 | 237  | 
assumes "Linear_order r"  | 
238  | 
and "a \<in> Field r"  | 
|
239  | 
and "b \<in> Field r"  | 
|
240  | 
shows "(a, b) \<in> r \<longleftrightarrow> (b, a) \<notin> r - Id"  | 
|
241  | 
using assms unfolding order_on_defs total_on_def antisym_def Id_def refl_on_def by force  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
242  | 
|
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
243  | 
|
| 60758 | 244  | 
subsubsection \<open>The upper and lower bounds operators\<close>  | 
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
245  | 
|
| 63572 | 246  | 
text \<open>  | 
247  | 
Here we define upper (``above") and lower (``below") bounds operators. We  | 
|
248  | 
think of \<open>r\<close> as a \<^emph>\<open>non-strict\<close> relation. The suffix \<open>S\<close> at the names of  | 
|
249  | 
some operators indicates that the bounds are strict -- e.g., \<open>underS a\<close> is  | 
|
250  | 
the set of all strict lower bounds of \<open>a\<close> (w.r.t. \<open>r\<close>). Capitalization of  | 
|
251  | 
the first letter in the name reminds that the operator acts on sets, rather  | 
|
252  | 
than on individual elements.  | 
|
253  | 
\<close>  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
254  | 
|
| 63572 | 255  | 
definition under :: "'a rel \<Rightarrow> 'a \<Rightarrow> 'a set"  | 
256  | 
  where "under r a \<equiv> {b. (b, a) \<in> r}"
 | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
257  | 
|
| 63572 | 258  | 
definition underS :: "'a rel \<Rightarrow> 'a \<Rightarrow> 'a set"  | 
259  | 
  where "underS r a \<equiv> {b. b \<noteq> a \<and> (b, a) \<in> r}"
 | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
260  | 
|
| 63572 | 261  | 
definition Under :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a set"  | 
262  | 
  where "Under r A \<equiv> {b \<in> Field r. \<forall>a \<in> A. (b, a) \<in> r}"
 | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
263  | 
|
| 63572 | 264  | 
definition UnderS :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a set"  | 
265  | 
  where "UnderS r A \<equiv> {b \<in> Field r. \<forall>a \<in> A. b \<noteq> a \<and> (b, a) \<in> r}"
 | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
266  | 
|
| 63572 | 267  | 
definition above :: "'a rel \<Rightarrow> 'a \<Rightarrow> 'a set"  | 
268  | 
  where "above r a \<equiv> {b. (a, b) \<in> r}"
 | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
269  | 
|
| 63572 | 270  | 
definition aboveS :: "'a rel \<Rightarrow> 'a \<Rightarrow> 'a set"  | 
271  | 
  where "aboveS r a \<equiv> {b. b \<noteq> a \<and> (a, b) \<in> r}"
 | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
272  | 
|
| 63572 | 273  | 
definition Above :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a set"  | 
274  | 
  where "Above r A \<equiv> {b \<in> Field r. \<forall>a \<in> A. (a, b) \<in> r}"
 | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
275  | 
|
| 63572 | 276  | 
definition AboveS :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a set"  | 
277  | 
  where "AboveS r A \<equiv> {b \<in> Field r. \<forall>a \<in> A. b \<noteq> a \<and> (a, b) \<in> r}"
 | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
278  | 
|
| 55173 | 279  | 
definition ofilter :: "'a rel \<Rightarrow> 'a set \<Rightarrow> bool"  | 
| 63572 | 280  | 
where "ofilter r A \<equiv> A \<subseteq> Field r \<and> (\<forall>a \<in> A. under r a \<subseteq> A)"  | 
| 55173 | 281  | 
|
| 63572 | 282  | 
text \<open>  | 
283  | 
Note: In the definitions of \<open>Above[S]\<close> and \<open>Under[S]\<close>, we bounded  | 
|
284  | 
comprehension by \<open>Field r\<close> in order to properly cover the case of \<open>A\<close> being  | 
|
285  | 
empty.  | 
|
286  | 
\<close>  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
287  | 
|
| 63572 | 288  | 
lemma underS_subset_under: "underS r a \<subseteq> under r a"  | 
289  | 
by (auto simp add: underS_def under_def)  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
290  | 
|
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
291  | 
lemma underS_notIn: "a \<notin> underS r a"  | 
| 63572 | 292  | 
by (simp add: underS_def)  | 
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
293  | 
|
| 63572 | 294  | 
lemma Refl_under_in: "Refl r \<Longrightarrow> a \<in> Field r \<Longrightarrow> a \<in> under r a"  | 
295  | 
by (simp add: refl_on_def under_def)  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
296  | 
|
| 63572 | 297  | 
lemma AboveS_disjoint: "A \<inter> (AboveS r A) = {}"
 | 
298  | 
by (auto simp add: AboveS_def)  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
299  | 
|
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
300  | 
lemma in_AboveS_underS: "a \<in> Field r \<Longrightarrow> a \<in> AboveS r (underS r a)"  | 
| 63572 | 301  | 
by (auto simp add: AboveS_def underS_def)  | 
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
302  | 
|
| 63572 | 303  | 
lemma Refl_under_underS: "Refl r \<Longrightarrow> a \<in> Field r \<Longrightarrow> under r a = underS r a \<union> {a}"
 | 
304  | 
unfolding under_def underS_def  | 
|
305  | 
using refl_on_def[of _ r] by fastforce  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
306  | 
|
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
307  | 
lemma underS_empty: "a \<notin> Field r \<Longrightarrow> underS r a = {}"
 | 
| 63572 | 308  | 
by (auto simp: Field_def underS_def)  | 
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
309  | 
|
| 63572 | 310  | 
lemma under_Field: "under r a \<subseteq> Field r"  | 
311  | 
by (auto simp: under_def Field_def)  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
312  | 
|
| 63572 | 313  | 
lemma underS_Field: "underS r a \<subseteq> Field r"  | 
314  | 
by (auto simp: underS_def Field_def)  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
315  | 
|
| 63572 | 316  | 
lemma underS_Field2: "a \<in> Field r \<Longrightarrow> underS r a \<subset> Field r"  | 
317  | 
using underS_notIn underS_Field by fast  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
318  | 
|
| 63572 | 319  | 
lemma underS_Field3: "Field r \<noteq> {} \<Longrightarrow> underS r a \<subset> Field r"
 | 
320  | 
by (cases "a \<in> Field r") (auto simp: underS_Field2 underS_empty)  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
321  | 
|
| 63572 | 322  | 
lemma AboveS_Field: "AboveS r A \<subseteq> Field r"  | 
323  | 
by (auto simp: AboveS_def Field_def)  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
324  | 
|
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
325  | 
lemma under_incr:  | 
| 63572 | 326  | 
assumes "trans r"  | 
327  | 
and "(a, b) \<in> r"  | 
|
328  | 
shows "under r a \<subseteq> under r b"  | 
|
329  | 
unfolding under_def  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
330  | 
proof safe  | 
| 63572 | 331  | 
fix x assume "(x, a) \<in> r"  | 
332  | 
with assms trans_def[of r] show "(x, b) \<in> r" by blast  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
333  | 
qed  | 
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
334  | 
|
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
335  | 
lemma underS_incr:  | 
| 63572 | 336  | 
assumes "trans r"  | 
337  | 
and "antisym r"  | 
|
338  | 
and ab: "(a, b) \<in> r"  | 
|
339  | 
shows "underS r a \<subseteq> underS r b"  | 
|
340  | 
unfolding underS_def  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
341  | 
proof safe  | 
| 63572 | 342  | 
assume *: "b \<noteq> a" and **: "(b, a) \<in> r"  | 
343  | 
with \<open>antisym r\<close> antisym_def[of r] ab show False  | 
|
344  | 
by blast  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
345  | 
next  | 
| 63572 | 346  | 
fix x assume "x \<noteq> a" "(x, a) \<in> r"  | 
347  | 
with ab \<open>trans r\<close> trans_def[of r] show "(x, b) \<in> r"  | 
|
348  | 
by blast  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
349  | 
qed  | 
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
350  | 
|
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
351  | 
lemma underS_incl_iff:  | 
| 63572 | 352  | 
assumes LO: "Linear_order r"  | 
353  | 
and INa: "a \<in> Field r"  | 
|
354  | 
and INb: "b \<in> Field r"  | 
|
355  | 
shows "underS r a \<subseteq> underS r b \<longleftrightarrow> (a, b) \<in> r"  | 
|
356  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
357  | 
proof  | 
| 63572 | 358  | 
assume ?rhs  | 
359  | 
with \<open>Linear_order r\<close> show ?lhs  | 
|
360  | 
by (simp add: order_on_defs underS_incr)  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
361  | 
next  | 
| 63572 | 362  | 
assume *: ?lhs  | 
363  | 
have "(a, b) \<in> r" if "a = b"  | 
|
364  | 
using assms that by (simp add: order_on_defs refl_on_def)  | 
|
365  | 
moreover have False if "a \<noteq> b" "(b, a) \<in> r"  | 
|
366  | 
proof -  | 
|
367  | 
from that have "b \<in> underS r a" unfolding underS_def by blast  | 
|
368  | 
with * have "b \<in> underS r b" by blast  | 
|
369  | 
then show ?thesis by (simp add: underS_notIn)  | 
|
370  | 
qed  | 
|
371  | 
ultimately show "(a,b) \<in> r"  | 
|
372  | 
using assms order_on_defs[of "Field r" r] total_on_def[of "Field r" r] by blast  | 
|
| 
55026
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
373  | 
qed  | 
| 
 
258fa7b5a621
folded 'Order_Relation_More_FP' into 'Order_Relation'
 
blanchet 
parents: 
54552 
diff
changeset
 | 
374  | 
|
| 70180 | 375  | 
lemma finite_Partial_order_induct[consumes 3, case_names step]:  | 
376  | 
assumes "Partial_order r"  | 
|
377  | 
and "x \<in> Field r"  | 
|
378  | 
and "finite r"  | 
|
379  | 
and step: "\<And>x. x \<in> Field r \<Longrightarrow> (\<And>y. y \<in> aboveS r x \<Longrightarrow> P y) \<Longrightarrow> P x"  | 
|
380  | 
shows "P x"  | 
|
381  | 
using assms(2)  | 
|
382  | 
proof (induct rule: wf_induct[of "r\<inverse> - Id"])  | 
|
383  | 
case 1  | 
|
384  | 
from assms(1,3) show "wf (r\<inverse> - Id)"  | 
|
385  | 
using partial_order_on_well_order_on partial_order_on_converse by blast  | 
|
386  | 
next  | 
|
387  | 
case prems: (2 x)  | 
|
388  | 
show ?case  | 
|
389  | 
by (rule step) (use prems in \<open>auto simp: aboveS_def intro: FieldI2\<close>)  | 
|
390  | 
qed  | 
|
391  | 
||
| 
63561
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
392  | 
lemma finite_Linear_order_induct[consumes 3, case_names step]:  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
393  | 
assumes "Linear_order r"  | 
| 63572 | 394  | 
and "x \<in> Field r"  | 
395  | 
and "finite r"  | 
|
396  | 
and step: "\<And>x. x \<in> Field r \<Longrightarrow> (\<And>y. y \<in> aboveS r x \<Longrightarrow> P y) \<Longrightarrow> P x"  | 
|
| 
63561
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
397  | 
shows "P x"  | 
| 63572 | 398  | 
using assms(2)  | 
399  | 
proof (induct rule: wf_induct[of "r\<inverse> - Id"])  | 
|
400  | 
case 1  | 
|
| 
63561
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
401  | 
from assms(1,3) show "wf (r\<inverse> - Id)"  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
402  | 
using linear_order_on_well_order_on linear_order_on_converse  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
403  | 
unfolding well_order_on_def by blast  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
404  | 
next  | 
| 63572 | 405  | 
case prems: (2 x)  | 
406  | 
show ?case  | 
|
407  | 
by (rule step) (use prems in \<open>auto simp: aboveS_def intro: FieldI2\<close>)  | 
|
| 
63561
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
408  | 
qed  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
61799 
diff
changeset
 | 
409  | 
|
| 55027 | 410  | 
|
| 60758 | 411  | 
subsection \<open>Variations on Well-Founded Relations\<close>  | 
| 55027 | 412  | 
|
| 60758 | 413  | 
text \<open>  | 
| 68484 | 414  | 
This subsection contains some variations of the results from \<^theory>\<open>HOL.Wellfounded\<close>:  | 
| 63572 | 415  | 
\<^item> means for slightly more direct definitions by well-founded recursion;  | 
416  | 
\<^item> variations of well-founded induction;  | 
|
417  | 
\<^item> means for proving a linear order to be a well-order.  | 
|
| 60758 | 418  | 
\<close>  | 
| 55027 | 419  | 
|
420  | 
||
| 60758 | 421  | 
subsubsection \<open>Characterizations of well-foundedness\<close>  | 
| 55027 | 422  | 
|
| 63572 | 423  | 
text \<open>  | 
424  | 
A transitive relation is well-founded iff it is ``locally'' well-founded,  | 
|
425  | 
i.e., iff its restriction to the lower bounds of of any element is  | 
|
426  | 
well-founded.  | 
|
427  | 
\<close>  | 
|
| 55027 | 428  | 
|
429  | 
lemma trans_wf_iff:  | 
|
| 63572 | 430  | 
assumes "trans r"  | 
431  | 
  shows "wf r \<longleftrightarrow> (\<forall>a. wf (r \<inter> (r\<inverse>``{a} \<times> r\<inverse>``{a})))"
 | 
|
432  | 
proof -  | 
|
433  | 
  define R where "R a = r \<inter> (r\<inverse>``{a} \<times> r\<inverse>``{a})" for a
 | 
|
434  | 
have "wf (R a)" if "wf r" for a  | 
|
435  | 
using that R_def wf_subset[of r "R a"] by auto  | 
|
| 55027 | 436  | 
moreover  | 
| 63572 | 437  | 
have "wf r" if *: "\<forall>a. wf(R a)"  | 
438  | 
unfolding wf_def  | 
|
439  | 
proof clarify  | 
|
440  | 
fix phi a  | 
|
441  | 
assume **: "\<forall>a. (\<forall>b. (b, a) \<in> r \<longrightarrow> phi b) \<longrightarrow> phi a"  | 
|
442  | 
define chi where "chi b \<longleftrightarrow> (b, a) \<in> r \<longrightarrow> phi b" for b  | 
|
443  | 
with * have "wf (R a)" by auto  | 
|
444  | 
then have "(\<forall>b. (\<forall>c. (c, b) \<in> R a \<longrightarrow> chi c) \<longrightarrow> chi b) \<longrightarrow> (\<forall>b. chi b)"  | 
|
445  | 
unfolding wf_def by blast  | 
|
446  | 
also have "\<forall>b. (\<forall>c. (c, b) \<in> R a \<longrightarrow> chi c) \<longrightarrow> chi b"  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
447  | 
proof safe  | 
| 63572 | 448  | 
fix b  | 
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
449  | 
assume "\<forall>c. (c, b) \<in> R a \<longrightarrow> chi c"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
450  | 
moreover have "(b, a) \<in> r \<Longrightarrow> \<forall>c. (c, b) \<in> r \<and> (c, a) \<in> r \<longrightarrow> phi c \<Longrightarrow> phi b"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
451  | 
proof -  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
452  | 
assume "(b, a) \<in> r" and "\<forall>c. (c, b) \<in> r \<and> (c, a) \<in> r \<longrightarrow> phi c"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
453  | 
then have "\<forall>c. (c, b) \<in> r \<longrightarrow> phi c"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
454  | 
using assms trans_def[of r] by blast  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
455  | 
with ** show "phi b" by blast  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
456  | 
qed  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
457  | 
ultimately show "chi b"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
458  | 
by (auto simp add: chi_def R_def)  | 
| 63572 | 459  | 
qed  | 
460  | 
finally have "\<forall>b. chi b" .  | 
|
461  | 
with ** chi_def show "phi a" by blast  | 
|
462  | 
qed  | 
|
463  | 
ultimately show ?thesis unfolding R_def by blast  | 
|
| 55027 | 464  | 
qed  | 
465  | 
||
| 
63952
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63572 
diff
changeset
 | 
466  | 
text\<open>A transitive relation is well-founded if all initial segments are finite.\<close>  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63572 
diff
changeset
 | 
467  | 
corollary wf_finite_segments:  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63572 
diff
changeset
 | 
468  | 
  assumes "irrefl r" and "trans r" and "\<And>x. finite {y. (y, x) \<in> r}"
 | 
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
469  | 
shows "wf r"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
470  | 
proof -  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
471  | 
  have "\<And>a. acyclic (r \<inter> {x. (x, a) \<in> r} \<times> {x. (x, a) \<in> r})"
 | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
472  | 
proof -  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
473  | 
fix a  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
474  | 
    have "trans (r \<inter> ({x. (x, a) \<in> r} \<times> {x. (x, a) \<in> r}))"
 | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
475  | 
using assms unfolding trans_def Field_def by blast  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
476  | 
    then show "acyclic (r \<inter> {x. (x, a) \<in> r} \<times> {x. (x, a) \<in> r})"
 | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
477  | 
using assms acyclic_def assms irrefl_def by fastforce  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
478  | 
qed  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
479  | 
then show ?thesis  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
480  | 
by (clarsimp simp: trans_wf_iff wf_iff_acyclic_if_finite converse_def assms)  | 
| 
63952
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63572 
diff
changeset
 | 
481  | 
qed  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63572 
diff
changeset
 | 
482  | 
|
| 61799 | 483  | 
text \<open>The next lemma is a variation of \<open>wf_eq_minimal\<close> from Wellfounded,  | 
| 63572 | 484  | 
allowing one to assume the set included in the field.\<close>  | 
| 55027 | 485  | 
|
| 63572 | 486  | 
lemma wf_eq_minimal2: "wf r \<longleftrightarrow> (\<forall>A. A \<subseteq> Field r \<and> A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a', a) \<notin> r))"
 | 
| 55027 | 487  | 
proof-  | 
| 63572 | 488  | 
  let ?phi = "\<lambda>A. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r)"
 | 
489  | 
have "wf r \<longleftrightarrow> (\<forall>A. ?phi A)"  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
490  | 
proof  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
491  | 
assume "wf r"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
492  | 
show "\<forall>A. ?phi A"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
493  | 
proof clarify  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
494  | 
fix A:: "'a set"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
495  | 
      assume "A \<noteq> {}"
 | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
496  | 
then obtain x where "x \<in> A"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
497  | 
by auto  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
498  | 
show "\<exists>a\<in>A. \<forall>a'\<in>A. (a', a) \<notin> r"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
499  | 
apply (rule wfE_min[of r x A])  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
500  | 
apply fact+  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
501  | 
by blast  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
502  | 
qed  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
503  | 
next  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
504  | 
assume *: "\<forall>A. ?phi A"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
505  | 
then show "wf r"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
506  | 
apply (clarsimp simp: ex_in_conv [THEN sym])  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
507  | 
apply (rule wfI_min)  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
508  | 
by fast  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
70180 
diff
changeset
 | 
509  | 
qed  | 
| 63572 | 510  | 
also have "(\<forall>A. ?phi A) \<longleftrightarrow> (\<forall>B \<subseteq> Field r. ?phi B)"  | 
| 55027 | 511  | 
proof  | 
512  | 
assume "\<forall>A. ?phi A"  | 
|
| 63572 | 513  | 
then show "\<forall>B \<subseteq> Field r. ?phi B" by simp  | 
| 55027 | 514  | 
next  | 
| 63572 | 515  | 
assume *: "\<forall>B \<subseteq> Field r. ?phi B"  | 
| 55027 | 516  | 
show "\<forall>A. ?phi A"  | 
| 63572 | 517  | 
proof clarify  | 
518  | 
fix A :: "'a set"  | 
|
519  | 
      assume **: "A \<noteq> {}"
 | 
|
520  | 
define B where "B = A \<inter> Field r"  | 
|
521  | 
show "\<exists>a \<in> A. \<forall>a' \<in> A. (a', a) \<notin> r"  | 
|
522  | 
      proof (cases "B = {}")
 | 
|
523  | 
case True  | 
|
524  | 
with ** obtain a where a: "a \<in> A" "a \<notin> Field r"  | 
|
525  | 
unfolding B_def by blast  | 
|
526  | 
with a have "\<forall>a' \<in> A. (a',a) \<notin> r"  | 
|
527  | 
unfolding Field_def by blast  | 
|
528  | 
with a show ?thesis by blast  | 
|
| 55027 | 529  | 
next  | 
| 63572 | 530  | 
case False  | 
531  | 
have "B \<subseteq> Field r" unfolding B_def by blast  | 
|
532  | 
with False * obtain a where a: "a \<in> B" "\<forall>a' \<in> B. (a', a) \<notin> r"  | 
|
533  | 
by blast  | 
|
534  | 
have "(a', a) \<notin> r" if "a' \<in> A" for a'  | 
|
535  | 
proof  | 
|
536  | 
assume a'a: "(a', a) \<in> r"  | 
|
537  | 
with that have "a' \<in> B" unfolding B_def Field_def by blast  | 
|
538  | 
with a a'a show False by blast  | 
|
| 55027 | 539  | 
qed  | 
| 63572 | 540  | 
with a show ?thesis unfolding B_def by blast  | 
| 55027 | 541  | 
qed  | 
542  | 
qed  | 
|
543  | 
qed  | 
|
544  | 
finally show ?thesis by blast  | 
|
545  | 
qed  | 
|
546  | 
||
547  | 
||
| 60758 | 548  | 
subsubsection \<open>Characterizations of well-foundedness\<close>  | 
| 55027 | 549  | 
|
| 63572 | 550  | 
text \<open>  | 
551  | 
The next lemma and its corollary enable one to prove that a linear order is  | 
|
552  | 
a well-order in a way which is more standard than via well-foundedness of  | 
|
553  | 
the strict version of the relation.  | 
|
554  | 
\<close>  | 
|
| 55027 | 555  | 
|
556  | 
lemma Linear_order_wf_diff_Id:  | 
|
| 63572 | 557  | 
assumes "Linear_order r"  | 
558  | 
  shows "wf (r - Id) \<longleftrightarrow> (\<forall>A \<subseteq> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r))"
 | 
|
559  | 
proof (cases "r \<subseteq> Id")  | 
|
560  | 
case True  | 
|
561  | 
  then have *: "r - Id = {}" by blast
 | 
|
562  | 
have "wf (r - Id)" by (simp add: *)  | 
|
563  | 
moreover have "\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r"  | 
|
564  | 
    if *: "A \<subseteq> Field r" and **: "A \<noteq> {}" for A
 | 
|
565  | 
proof -  | 
|
566  | 
from \<open>Linear_order r\<close> True  | 
|
567  | 
    obtain a where a: "r = {} \<or> r = {(a, a)}"
 | 
|
568  | 
unfolding order_on_defs using Total_subset_Id [of r] by blast  | 
|
569  | 
    with * ** have "A = {a} \<and> r = {(a, a)}"
 | 
|
570  | 
unfolding Field_def by blast  | 
|
571  | 
with a show ?thesis by blast  | 
|
572  | 
qed  | 
|
| 55027 | 573  | 
ultimately show ?thesis by blast  | 
574  | 
next  | 
|
| 63572 | 575  | 
case False  | 
576  | 
with \<open>Linear_order r\<close> have Field: "Field r = Field (r - Id)"  | 
|
577  | 
unfolding order_on_defs using Total_Id_Field [of r] by blast  | 
|
| 55027 | 578  | 
show ?thesis  | 
579  | 
proof  | 
|
| 63572 | 580  | 
assume *: "wf (r - Id)"  | 
581  | 
    show "\<forall>A \<subseteq> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r)"
 | 
|
582  | 
proof clarify  | 
|
583  | 
fix A  | 
|
584  | 
      assume **: "A \<subseteq> Field r" and ***: "A \<noteq> {}"
 | 
|
585  | 
then have "\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r - Id"  | 
|
586  | 
using Field * unfolding wf_eq_minimal2 by simp  | 
|
587  | 
moreover have "\<forall>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r \<longleftrightarrow> (a', a) \<notin> r - Id"  | 
|
588  | 
using Linear_order_in_diff_Id [OF \<open>Linear_order r\<close>] ** by blast  | 
|
589  | 
ultimately show "\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r" by blast  | 
|
| 55027 | 590  | 
qed  | 
591  | 
next  | 
|
| 63572 | 592  | 
    assume *: "\<forall>A \<subseteq> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r)"
 | 
593  | 
show "wf (r - Id)"  | 
|
594  | 
unfolding wf_eq_minimal2  | 
|
595  | 
proof clarify  | 
|
596  | 
fix A  | 
|
597  | 
      assume **: "A \<subseteq> Field(r - Id)" and ***: "A \<noteq> {}"
 | 
|
598  | 
then have "\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r"  | 
|
599  | 
using Field * by simp  | 
|
600  | 
moreover have "\<forall>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r \<longleftrightarrow> (a', a) \<notin> r - Id"  | 
|
601  | 
using Linear_order_in_diff_Id [OF \<open>Linear_order r\<close>] ** mono_Field[of "r - Id" r] by blast  | 
|
602  | 
ultimately show "\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r - Id"  | 
|
603  | 
by blast  | 
|
| 55027 | 604  | 
qed  | 
605  | 
qed  | 
|
606  | 
qed  | 
|
607  | 
||
608  | 
corollary Linear_order_Well_order_iff:  | 
|
| 63572 | 609  | 
"Linear_order r \<Longrightarrow>  | 
610  | 
    Well_order r \<longleftrightarrow> (\<forall>A \<subseteq> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a, a') \<in> r))"
 | 
|
611  | 
unfolding well_order_on_def using Linear_order_wf_diff_Id[of r] by blast  | 
|
| 55027 | 612  | 
|
| 26273 | 613  | 
end  |