author | wenzelm |
Mon, 09 Dec 2013 12:22:23 +0100 | |
changeset 54703 | 499f92dc6e45 |
parent 54552 | 5d57cbec0f0f |
child 55026 | 258fa7b5a621 |
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 |
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
3 |
*) |
26273 | 4 |
|
5 |
header {* Orders as Relations *} |
|
6 |
||
7 |
theory Order_Relation |
|
54552
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
8 |
imports Wellfounded |
26273 | 9 |
begin |
10 |
||
26295 | 11 |
subsection{* Orders on a set *} |
12 |
||
30198 | 13 |
definition "preorder_on A r \<equiv> refl_on A r \<and> trans r" |
26295 | 14 |
|
15 |
definition "partial_order_on A r \<equiv> preorder_on A r \<and> antisym r" |
|
26273 | 16 |
|
26295 | 17 |
definition "linear_order_on A r \<equiv> partial_order_on A r \<and> total_on A r" |
18 |
||
19 |
definition "strict_linear_order_on A r \<equiv> trans r \<and> irrefl r \<and> total_on A r" |
|
20 |
||
21 |
definition "well_order_on A r \<equiv> linear_order_on A r \<and> wf(r - Id)" |
|
26273 | 22 |
|
26295 | 23 |
lemmas order_on_defs = |
24 |
preorder_on_def partial_order_on_def linear_order_on_def |
|
25 |
strict_linear_order_on_def well_order_on_def |
|
26 |
||
26273 | 27 |
|
26295 | 28 |
lemma preorder_on_empty[simp]: "preorder_on {} {}" |
29 |
by(simp add:preorder_on_def trans_def) |
|
30 |
||
31 |
lemma partial_order_on_empty[simp]: "partial_order_on {} {}" |
|
32 |
by(simp add:partial_order_on_def) |
|
26273 | 33 |
|
26295 | 34 |
lemma lnear_order_on_empty[simp]: "linear_order_on {} {}" |
35 |
by(simp add:linear_order_on_def) |
|
36 |
||
37 |
lemma well_order_on_empty[simp]: "well_order_on {} {}" |
|
38 |
by(simp add:well_order_on_def) |
|
39 |
||
26273 | 40 |
|
26295 | 41 |
lemma preorder_on_converse[simp]: "preorder_on A (r^-1) = preorder_on A r" |
42 |
by (simp add:preorder_on_def) |
|
43 |
||
44 |
lemma partial_order_on_converse[simp]: |
|
45 |
"partial_order_on A (r^-1) = partial_order_on A r" |
|
46 |
by (simp add: partial_order_on_def) |
|
26273 | 47 |
|
26295 | 48 |
lemma linear_order_on_converse[simp]: |
49 |
"linear_order_on A (r^-1) = linear_order_on A r" |
|
50 |
by (simp add: linear_order_on_def) |
|
51 |
||
26273 | 52 |
|
26295 | 53 |
lemma strict_linear_order_on_diff_Id: |
54 |
"linear_order_on A r \<Longrightarrow> strict_linear_order_on A (r-Id)" |
|
55 |
by(simp add: order_on_defs trans_diff_Id) |
|
56 |
||
57 |
||
58 |
subsection{* Orders on the field *} |
|
26273 | 59 |
|
30198 | 60 |
abbreviation "Refl r \<equiv> refl_on (Field r) r" |
26295 | 61 |
|
62 |
abbreviation "Preorder r \<equiv> preorder_on (Field r) r" |
|
63 |
||
64 |
abbreviation "Partial_order r \<equiv> partial_order_on (Field r) r" |
|
26273 | 65 |
|
26295 | 66 |
abbreviation "Total r \<equiv> total_on (Field r) r" |
67 |
||
68 |
abbreviation "Linear_order r \<equiv> linear_order_on (Field r) r" |
|
69 |
||
70 |
abbreviation "Well_order r \<equiv> well_order_on (Field r) r" |
|
71 |
||
26273 | 72 |
|
73 |
lemma subset_Image_Image_iff: |
|
74 |
"\<lbrakk> Preorder r; A \<subseteq> Field r; B \<subseteq> Field r\<rbrakk> \<Longrightarrow> |
|
75 |
r `` A \<subseteq> r `` B \<longleftrightarrow> (\<forall>a\<in>A.\<exists>b\<in>B. (b,a):r)" |
|
48750 | 76 |
unfolding preorder_on_def refl_on_def Image_def |
77 |
apply (simp add: subset_eq) |
|
78 |
unfolding trans_def by fast |
|
26273 | 79 |
|
80 |
lemma subset_Image1_Image1_iff: |
|
81 |
"\<lbrakk> Preorder r; a : Field r; b : Field r\<rbrakk> \<Longrightarrow> r `` {a} \<subseteq> r `` {b} \<longleftrightarrow> (b,a):r" |
|
82 |
by(simp add:subset_Image_Image_iff) |
|
83 |
||
84 |
lemma Refl_antisym_eq_Image1_Image1_iff: |
|
54552
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
85 |
assumes r: "Refl r" and as: "antisym r" and abf: "a \<in> Field r" "b \<in> Field r" |
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
86 |
shows "r `` {a} = r `` {b} \<longleftrightarrow> a = b" |
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
87 |
proof |
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
88 |
assume "r `` {a} = r `` {b}" |
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
89 |
hence e: "\<And>x. (a, x) \<in> r \<longleftrightarrow> (b, x) \<in> r" by (simp add: set_eq_iff) |
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
90 |
have "(a, a) \<in> r" "(b, b) \<in> r" using r abf by (simp_all add: refl_on_def) |
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
91 |
hence "(a, b) \<in> r" "(b, a) \<in> r" using e[of a] e[of b] by simp_all |
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
92 |
thus "a = b" using as[unfolded antisym_def] by blast |
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
93 |
qed fast |
26273 | 94 |
|
95 |
lemma Partial_order_eq_Image1_Image1_iff: |
|
96 |
"\<lbrakk>Partial_order r; a:Field r; b:Field r\<rbrakk> \<Longrightarrow> r `` {a} = r `` {b} \<longleftrightarrow> a=b" |
|
26295 | 97 |
by(auto simp:order_on_defs Refl_antisym_eq_Image1_Image1_iff) |
98 |
||
52182 | 99 |
lemma Total_Id_Field: |
100 |
assumes TOT: "Total r" and NID: "\<not> (r <= Id)" |
|
101 |
shows "Field r = Field(r - Id)" |
|
102 |
using mono_Field[of "r - Id" r] Diff_subset[of r Id] |
|
103 |
proof(auto) |
|
104 |
have "r \<noteq> {}" using NID by fast |
|
54482 | 105 |
then obtain b and c where "b \<noteq> c \<and> (b,c) \<in> r" using NID by auto |
52182 | 106 |
hence 1: "b \<noteq> c \<and> {b,c} \<le> Field r" by (auto simp: Field_def) |
54552
5d57cbec0f0f
moving 'Order_Relation' to 'HOL' (since it's a BNF dependency)
blanchet
parents:
54551
diff
changeset
|
107 |
|
52182 | 108 |
fix a assume *: "a \<in> Field r" |
109 |
obtain d where 2: "d \<in> Field r" and 3: "d \<noteq> a" |
|
110 |
using * 1 by auto |
|
111 |
hence "(a,d) \<in> r \<or> (d,a) \<in> r" using * TOT |
|
112 |
by (simp add: total_on_def) |
|
113 |
thus "a \<in> Field(r - Id)" using 3 unfolding Field_def by blast |
|
114 |
qed |
|
115 |
||
26295 | 116 |
|
117 |
subsection{* Orders on a type *} |
|
118 |
||
119 |
abbreviation "strict_linear_order \<equiv> strict_linear_order_on UNIV" |
|
120 |
||
121 |
abbreviation "linear_order \<equiv> linear_order_on UNIV" |
|
122 |
||
54551 | 123 |
abbreviation "well_order \<equiv> well_order_on UNIV" |
26273 | 124 |
|
125 |
end |