| author | wenzelm |
| Wed, 15 Nov 2006 20:50:22 +0100 | |
| changeset 21391 | a8809f46bd7f |
| parent 21383 | 17e6275e13f5 |
| child 21404 | eb85850d3eb7 |
| permissions | -rw-r--r-- |
| 15524 | 1 |
(* Title: HOL/Orderings.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson |
|
4 |
*) |
|
5 |
||
| 21329 | 6 |
header {* Syntactic and abstract orders *}
|
| 15524 | 7 |
|
8 |
theory Orderings |
|
| 21329 | 9 |
imports HOL |
| 15524 | 10 |
begin |
11 |
||
| 21329 | 12 |
section {* Abstract orders *}
|
| 21083 | 13 |
|
| 21329 | 14 |
subsection {* Order syntax *}
|
| 15524 | 15 |
|
| 21194 | 16 |
class ord = |
| 20588 | 17 |
fixes less_eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
| 21204 | 18 |
and less :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
19 |
begin |
|
20 |
||
21 |
notation |
|
22 |
less_eq ("op \<^loc><=")
|
|
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
23 |
less_eq ("(_/ \<^loc><= _)" [51, 51] 50)
|
| 21204 | 24 |
less ("op \<^loc><")
|
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
25 |
less ("(_/ \<^loc>< _)" [51, 51] 50)
|
| 21204 | 26 |
|
27 |
notation (xsymbols) |
|
28 |
less_eq ("op \<^loc>\<le>")
|
|
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
29 |
less_eq ("(_/ \<^loc>\<le> _)" [51, 51] 50)
|
| 15524 | 30 |
|
| 21204 | 31 |
notation (HTML output) |
32 |
less_eq ("op \<^loc>\<le>")
|
|
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
33 |
less_eq ("(_/ \<^loc>\<le> _)" [51, 51] 50)
|
| 21204 | 34 |
|
35 |
abbreviation (input) |
|
36 |
greater (infix "\<^loc>>" 50) |
|
37 |
"x \<^loc>> y \<equiv> y \<^loc>< x" |
|
38 |
greater_eq (infix "\<^loc>>=" 50) |
|
39 |
"x \<^loc>>= y \<equiv> y \<^loc><= x" |
|
40 |
||
41 |
notation (xsymbols) |
|
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
42 |
greater_eq (infix "\<^loc>\<ge>" 50) |
| 21204 | 43 |
|
44 |
end |
|
45 |
||
46 |
notation |
|
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
47 |
less_eq ("op <=")
|
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
48 |
less_eq ("(_/ <= _)" [51, 51] 50)
|
| 21204 | 49 |
less ("op <")
|
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
50 |
less ("(_/ < _)" [51, 51] 50)
|
| 21204 | 51 |
|
52 |
notation (xsymbols) |
|
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19637
diff
changeset
|
53 |
less_eq ("op \<le>")
|
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
54 |
less_eq ("(_/ \<le> _)" [51, 51] 50)
|
| 15524 | 55 |
|
| 21204 | 56 |
notation (HTML output) |
57 |
less_eq ("op \<le>")
|
|
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
58 |
less_eq ("(_/ \<le> _)" [51, 51] 50)
|
| 20714 | 59 |
|
| 19536 | 60 |
abbreviation (input) |
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
61 |
greater (infix ">" 50) |
| 20714 | 62 |
"x > y \<equiv> y < x" |
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
63 |
greater_eq (infix ">=" 50) |
| 20714 | 64 |
"x >= y \<equiv> y <= x" |
65 |
||
| 21204 | 66 |
notation (xsymbols) |
|
21259
63ab016c99ca
modified less/less_eq syntax to avoid "x < y < z" etc.;
wenzelm
parents:
21248
diff
changeset
|
67 |
greater_eq (infix "\<ge>" 50) |
| 15524 | 68 |
|
69 |
||
| 21329 | 70 |
subsection {* Quasiorders (preorders) *}
|
| 15524 | 71 |
|
| 21329 | 72 |
locale preorder = |
|
21383
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
73 |
fixes less_eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sqsubseteq>" 50) |
|
21216
1c8580913738
made locale partial_order compatible with axclass order; changed import order; consecutive changes
haftmann
parents:
21204
diff
changeset
|
74 |
fixes less :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sqsubset>" 50) |
|
1c8580913738
made locale partial_order compatible with axclass order; changed import order; consecutive changes
haftmann
parents:
21204
diff
changeset
|
75 |
assumes refl [iff]: "x \<sqsubseteq> x" |
|
1c8580913738
made locale partial_order compatible with axclass order; changed import order; consecutive changes
haftmann
parents:
21204
diff
changeset
|
76 |
and trans: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z" |
| 21248 | 77 |
and less_le: "x \<sqsubset> y \<longleftrightarrow> x \<sqsubseteq> y \<and> x \<noteq> y" |
78 |
begin |
|
79 |
||
80 |
abbreviation (input) |
|
81 |
greater_eq (infixl "\<sqsupseteq>" 50) |
|
82 |
"x \<sqsupseteq> y \<equiv> y \<sqsubseteq> x" |
|
83 |
||
84 |
abbreviation (input) |
|
85 |
greater (infixl "\<sqsupset>" 50) |
|
86 |
"x \<sqsupset> y \<equiv> y \<sqsubset> x" |
|
87 |
||
| 15524 | 88 |
text {* Reflexivity. *}
|
89 |
||
| 21248 | 90 |
lemma eq_refl: "x = y \<Longrightarrow> x \<sqsubseteq> y" |
| 15524 | 91 |
-- {* This form is useful with the classical reasoner. *}
|
| 21248 | 92 |
by (erule ssubst) (rule refl) |
| 15524 | 93 |
|
| 21248 | 94 |
lemma less_irrefl [iff]: "\<not> x \<sqsubset> x" |
95 |
by (simp add: less_le) |
|
| 15524 | 96 |
|
| 21248 | 97 |
lemma le_less: "x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubset> y \<or> x = y" |
| 15524 | 98 |
-- {* NOT suitable for iff, since it can cause PROOF FAILED. *}
|
| 21248 | 99 |
by (simp add: less_le) blast |
| 15524 | 100 |
|
| 21248 | 101 |
lemma le_imp_less_or_eq: "x \<sqsubseteq> y \<Longrightarrow> x \<sqsubset> y \<or> x = y" |
102 |
unfolding less_le by blast |
|
| 15524 | 103 |
|
| 21248 | 104 |
lemma less_imp_le: "x \<sqsubset> y \<Longrightarrow> x \<sqsubseteq> y" |
105 |
unfolding less_le by blast |
|
106 |
||
| 21329 | 107 |
lemma less_imp_neq: "x \<sqsubset> y \<Longrightarrow> x \<noteq> y" |
108 |
by (erule contrapos_pn, erule subst, rule less_irrefl) |
|
109 |
||
110 |
||
111 |
text {* Useful for simplification, but too risky to include by default. *}
|
|
112 |
||
113 |
lemma less_imp_not_eq: "x \<sqsubset> y \<Longrightarrow> (x = y) \<longleftrightarrow> False" |
|
114 |
by auto |
|
115 |
||
116 |
lemma less_imp_not_eq2: "x \<sqsubset> y \<Longrightarrow> (y = x) \<longleftrightarrow> False" |
|
117 |
by auto |
|
118 |
||
119 |
||
120 |
text {* Transitivity rules for calculational reasoning *}
|
|
121 |
||
122 |
lemma neq_le_trans: "\<lbrakk> a \<noteq> b; a \<sqsubseteq> b \<rbrakk> \<Longrightarrow> a \<sqsubset> b" |
|
123 |
by (simp add: less_le) |
|
124 |
||
125 |
lemma le_neq_trans: "\<lbrakk> a \<sqsubseteq> b; a \<noteq> b \<rbrakk> \<Longrightarrow> a \<sqsubset> b" |
|
126 |
by (simp add: less_le) |
|
127 |
||
128 |
end |
|
129 |
||
130 |
||
131 |
subsection {* Partial orderings *}
|
|
132 |
||
133 |
locale partial_order = preorder + |
|
134 |
assumes antisym: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y" |
|
135 |
||
136 |
context partial_order |
|
137 |
begin |
|
| 15524 | 138 |
|
139 |
text {* Asymmetry. *}
|
|
140 |
||
| 21248 | 141 |
lemma less_not_sym: "x \<sqsubset> y \<Longrightarrow> \<not> (y \<sqsubset> x)" |
142 |
by (simp add: less_le antisym) |
|
| 15524 | 143 |
|
| 21248 | 144 |
lemma less_asym: "x \<sqsubset> y \<Longrightarrow> (\<not> P \<Longrightarrow> y \<sqsubset> x) \<Longrightarrow> P" |
145 |
by (drule less_not_sym, erule contrapos_np) simp |
|
| 15524 | 146 |
|
| 21248 | 147 |
lemma eq_iff: "x = y \<longleftrightarrow> x \<sqsubseteq> y \<and> y \<sqsubseteq> x" |
148 |
by (blast intro: antisym) |
|
| 15524 | 149 |
|
| 21248 | 150 |
lemma antisym_conv: "y \<sqsubseteq> x \<Longrightarrow> x \<sqsubseteq> y \<longleftrightarrow> x = y" |
151 |
by (blast intro: antisym) |
|
| 15524 | 152 |
|
| 21248 | 153 |
lemma less_imp_neq: "x \<sqsubset> y \<Longrightarrow> x \<noteq> y" |
154 |
by (erule contrapos_pn, erule subst, rule less_irrefl) |
|
155 |
||
| 21083 | 156 |
|
| 15524 | 157 |
text {* Transitivity. *}
|
158 |
||
| 21248 | 159 |
lemma less_trans: "\<lbrakk> x \<sqsubset> y; y \<sqsubset> z \<rbrakk> \<Longrightarrow> x \<sqsubset> z" |
160 |
by (simp add: less_le) (blast intro: trans antisym) |
|
| 15524 | 161 |
|
| 21248 | 162 |
lemma le_less_trans: "\<lbrakk> x \<sqsubseteq> y; y \<sqsubset> z \<rbrakk> \<Longrightarrow> x \<sqsubset> z" |
163 |
by (simp add: less_le) (blast intro: trans antisym) |
|
| 15524 | 164 |
|
| 21248 | 165 |
lemma less_le_trans: "\<lbrakk> x \<sqsubset> y; y \<sqsubseteq> z \<rbrakk> \<Longrightarrow> x \<sqsubset> z" |
166 |
by (simp add: less_le) (blast intro: trans antisym) |
|
| 15524 | 167 |
|
168 |
||
169 |
text {* Useful for simplification, but too risky to include by default. *}
|
|
170 |
||
| 21248 | 171 |
lemma less_imp_not_less: "x \<sqsubset> y \<Longrightarrow> (\<not> y \<sqsubset> x) \<longleftrightarrow> True" |
172 |
by (blast elim: less_asym) |
|
| 15524 | 173 |
|
| 21248 | 174 |
lemma less_imp_triv: "x \<sqsubset> y \<Longrightarrow> (y \<sqsubset> x \<longrightarrow> P) \<longleftrightarrow> True" |
175 |
by (blast elim: less_asym) |
|
| 15524 | 176 |
|
| 21248 | 177 |
|
| 21083 | 178 |
text {* Transitivity rules for calculational reasoning *}
|
| 15524 | 179 |
|
| 21248 | 180 |
lemma less_asym': "\<lbrakk> a \<sqsubset> b; b \<sqsubset> a \<rbrakk> \<Longrightarrow> P" |
181 |
by (rule less_asym) |
|
182 |
||
183 |
end |
|
| 15524 | 184 |
|
| 21329 | 185 |
axclass order < ord |
186 |
order_refl [iff]: "x <= x" |
|
187 |
order_trans: "x <= y ==> y <= z ==> x <= z" |
|
188 |
order_antisym: "x <= y ==> y <= x ==> x = y" |
|
189 |
order_less_le: "(x < y) = (x <= y & x ~= y)" |
|
| 15524 | 190 |
|
| 21329 | 191 |
interpretation order: |
192 |
partial_order ["op \<le> \<Colon> 'a\<Colon>order \<Rightarrow> 'a \<Rightarrow> bool" "op < \<Colon> 'a\<Colon>order \<Rightarrow> 'a \<Rightarrow> bool"] |
|
193 |
apply unfold_locales |
|
194 |
apply (rule order_refl) |
|
195 |
apply (erule (1) order_trans) |
|
196 |
apply (rule order_less_le) |
|
197 |
apply (erule (1) order_antisym) |
|
198 |
done |
|
|
21216
1c8580913738
made locale partial_order compatible with axclass order; changed import order; consecutive changes
haftmann
parents:
21204
diff
changeset
|
199 |
|
| 21329 | 200 |
|
201 |
subsection {* Linear (total) orders *}
|
|
202 |
||
203 |
locale linorder = partial_order + |
|
|
21216
1c8580913738
made locale partial_order compatible with axclass order; changed import order; consecutive changes
haftmann
parents:
21204
diff
changeset
|
204 |
assumes linear: "x \<sqsubseteq> y \<or> y \<sqsubseteq> x" |
| 15524 | 205 |
|
| 21329 | 206 |
context linorder |
| 21248 | 207 |
begin |
208 |
||
209 |
lemma trichotomy: "x \<sqsubset> y \<or> x = y \<or> y \<sqsubset> x" |
|
210 |
unfolding less_le using less_le linear by blast |
|
211 |
||
212 |
lemma le_less_linear: "x \<sqsubseteq> y \<or> y \<sqsubset> x" |
|
213 |
by (simp add: le_less trichotomy) |
|
214 |
||
215 |
lemma le_cases [case_names le ge]: |
|
216 |
"\<lbrakk> x \<sqsubseteq> y \<Longrightarrow> P; y \<sqsubseteq> x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
|
217 |
using linear by blast |
|
218 |
||
219 |
lemma cases [case_names less equal greater]: |
|
220 |
"\<lbrakk> x \<sqsubset> y \<Longrightarrow> P; x = y \<Longrightarrow> P; y \<sqsubset> x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
|
221 |
using trichotomy by blast |
|
222 |
||
223 |
lemma not_less: "\<not> x \<sqsubset> y \<longleftrightarrow> y \<sqsubseteq> x" |
|
224 |
apply (simp add: less_le) |
|
225 |
using linear apply (blast intro: antisym) |
|
| 15524 | 226 |
done |
227 |
||
| 21248 | 228 |
lemma not_le: "\<not> x \<sqsubseteq> y \<longleftrightarrow> y \<sqsubset> x" |
229 |
apply (simp add: less_le) |
|
230 |
using linear apply (blast intro: antisym) |
|
| 15524 | 231 |
done |
232 |
||
| 21248 | 233 |
lemma neq_iff: "x \<noteq> y \<longleftrightarrow> x \<sqsubset> y \<or> y \<sqsubset> x" |
234 |
by (cut_tac x = x and y = y in trichotomy, auto) |
|
| 15524 | 235 |
|
| 21248 | 236 |
lemma neqE: "\<lbrakk> x \<noteq> y; x \<sqsubset> y \<Longrightarrow> R; y \<sqsubset> x \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
237 |
by (simp add: neq_iff) blast |
|
| 15524 | 238 |
|
| 21248 | 239 |
lemma antisym_conv1: "\<not> x \<sqsubset> y \<Longrightarrow> x \<sqsubseteq> y \<longleftrightarrow> x = y" |
240 |
by (blast intro: antisym dest: not_less [THEN iffD1]) |
|
| 15524 | 241 |
|
| 21248 | 242 |
lemma antisym_conv2: "x \<sqsubseteq> y \<Longrightarrow> \<not> x \<sqsubset> y \<longleftrightarrow> x = y" |
243 |
by (blast intro: antisym dest: not_less [THEN iffD1]) |
|
| 15524 | 244 |
|
| 21248 | 245 |
lemma antisym_conv3: "\<not> y \<sqsubset> x \<Longrightarrow> \<not> x \<sqsubset> y \<longleftrightarrow> x = y" |
246 |
by (blast intro: antisym dest: not_less [THEN iffD1]) |
|
| 15524 | 247 |
|
| 16796 | 248 |
text{*Replacing the old Nat.leI*}
|
| 21248 | 249 |
lemma leI: "\<not> x \<sqsubset> y \<Longrightarrow> y \<sqsubseteq> x" |
250 |
unfolding not_less . |
|
| 16796 | 251 |
|
| 21248 | 252 |
lemma leD: "y \<sqsubseteq> x \<Longrightarrow> \<not> x \<sqsubset> y" |
253 |
unfolding not_less . |
|
| 16796 | 254 |
|
255 |
(*FIXME inappropriate name (or delete altogether)*) |
|
| 21248 | 256 |
lemma not_leE: "\<not> y \<sqsubseteq> x \<Longrightarrow> x \<sqsubset> y" |
257 |
unfolding not_le . |
|
258 |
||
|
21383
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
259 |
(* min/max *) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
260 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
261 |
definition |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
262 |
min :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
263 |
"min a b = (if a \<sqsubseteq> b then a else b)" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
264 |
max :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
265 |
"max a b = (if a \<sqsubseteq> b then b else a)" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
266 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
267 |
lemma min_le_iff_disj: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
268 |
"min x y \<sqsubseteq> z \<longleftrightarrow> x \<sqsubseteq> z \<or> y \<sqsubseteq> z" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
269 |
unfolding min_def using linear by (auto intro: trans) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
270 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
271 |
lemma le_max_iff_disj: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
272 |
"z \<sqsubseteq> max x y \<longleftrightarrow> z \<sqsubseteq> x \<or> z \<sqsubseteq> y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
273 |
unfolding max_def using linear by (auto intro: trans) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
274 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
275 |
lemma min_less_iff_disj: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
276 |
"min x y \<sqsubset> z \<longleftrightarrow> x \<sqsubset> z \<or> y \<sqsubset> z" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
277 |
unfolding min_def le_less using trichotomy by (auto intro: less_trans) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
278 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
279 |
lemma less_max_iff_disj: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
280 |
"z \<sqsubset> max x y \<longleftrightarrow> z \<sqsubset> x \<or> z \<sqsubset> y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
281 |
unfolding max_def le_less using trichotomy by (auto intro: less_trans) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
282 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
283 |
lemma min_less_iff_conj [simp]: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
284 |
"z \<sqsubset> min x y \<longleftrightarrow> z \<sqsubset> x \<and> z \<sqsubset> y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
285 |
unfolding min_def le_less using trichotomy by (auto intro: less_trans) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
286 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
287 |
lemma max_less_iff_conj [simp]: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
288 |
"max x y \<sqsubset> z \<longleftrightarrow> x \<sqsubset> z \<and> y \<sqsubset> z" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
289 |
unfolding max_def le_less using trichotomy by (auto intro: less_trans) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
290 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
291 |
lemma split_min: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
292 |
"P (min i j) \<longleftrightarrow> (i \<sqsubseteq> j \<longrightarrow> P i) \<and> (\<not> i \<sqsubseteq> j \<longrightarrow> P j)" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
293 |
by (simp add: min_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
294 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
295 |
lemma split_max: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
296 |
"P (max i j) \<longleftrightarrow> (i \<sqsubseteq> j \<longrightarrow> P j) \<and> (\<not> i \<sqsubseteq> j \<longrightarrow> P i)" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
297 |
by (simp add: max_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
298 |
|
| 21248 | 299 |
end |
300 |
||
| 21329 | 301 |
axclass linorder < order |
302 |
linorder_linear: "x <= y | y <= x" |
|
303 |
||
304 |
interpretation linorder: |
|
305 |
linorder ["op \<le> \<Colon> 'a\<Colon>linorder \<Rightarrow> 'a \<Rightarrow> bool" "op < \<Colon> 'a\<Colon>linorder \<Rightarrow> 'a \<Rightarrow> bool"] |
|
306 |
by unfold_locales (rule linorder_linear) |
|
307 |
||
| 21248 | 308 |
|
309 |
subsection {* Name duplicates *}
|
|
310 |
||
311 |
lemmas order_eq_refl [where 'b = "?'a::order"] = order.eq_refl |
|
312 |
lemmas order_less_irrefl [where 'b = "?'a::order"] = order.less_irrefl |
|
313 |
lemmas order_le_less [where 'b = "?'a::order"] = order.le_less |
|
314 |
lemmas order_le_imp_less_or_eq [where 'b = "?'a::order"] = order.le_imp_less_or_eq |
|
315 |
lemmas order_less_imp_le [where 'b = "?'a::order"] = order.less_imp_le |
|
316 |
lemmas order_less_not_sym [where 'b = "?'a::order"] = order.less_not_sym |
|
317 |
lemmas order_less_asym [where 'b = "?'a::order"] = order.less_asym |
|
318 |
lemmas order_eq_iff [where 'b = "?'a::order"] = order.eq_iff |
|
319 |
lemmas order_antisym_conv [where 'b = "?'a::order"] = order.antisym_conv |
|
320 |
lemmas less_imp_neq [where 'b = "?'a::order"] = order.less_imp_neq |
|
321 |
lemmas order_less_trans [where 'b = "?'a::order"] = order.less_trans |
|
322 |
lemmas order_le_less_trans [where 'b = "?'a::order"] = order.le_less_trans |
|
323 |
lemmas order_less_le_trans [where 'b = "?'a::order"] = order.less_le_trans |
|
324 |
lemmas order_less_imp_not_less [where 'b = "?'a::order"] = order.less_imp_not_less |
|
325 |
lemmas order_less_imp_triv [where 'b = "?'a::order"] = order.less_imp_triv |
|
326 |
lemmas order_less_imp_not_eq [where 'b = "?'a::order"] = order.less_imp_not_eq |
|
327 |
lemmas order_less_imp_not_eq2 [where 'b = "?'a::order"] = order.less_imp_not_eq2 |
|
328 |
lemmas order_neq_le_trans [where 'b = "?'a::order"] = order.neq_le_trans |
|
329 |
lemmas order_le_neq_trans [where 'b = "?'a::order"] = order.le_neq_trans |
|
330 |
lemmas order_less_asym' [where 'b = "?'a::order"] = order.less_asym' |
|
331 |
lemmas linorder_less_linear [where 'b = "?'a::linorder"] = linorder.trichotomy |
|
332 |
lemmas linorder_le_less_linear [where 'b = "?'a::linorder"] = linorder.le_less_linear |
|
333 |
lemmas linorder_le_cases [where 'b = "?'a::linorder"] = linorder.le_cases |
|
334 |
lemmas linorder_cases [where 'b = "?'a::linorder"] = linorder.cases |
|
335 |
lemmas linorder_not_less [where 'b = "?'a::linorder"] = linorder.not_less |
|
336 |
lemmas linorder_not_le [where 'b = "?'a::linorder"] = linorder.not_le |
|
337 |
lemmas linorder_neq_iff [where 'b = "?'a::linorder"] = linorder.neq_iff |
|
338 |
lemmas linorder_neqE [where 'b = "?'a::linorder"] = linorder.neqE |
|
339 |
lemmas linorder_antisym_conv1 [where 'b = "?'a::linorder"] = linorder.antisym_conv1 |
|
340 |
lemmas linorder_antisym_conv2 [where 'b = "?'a::linorder"] = linorder.antisym_conv2 |
|
341 |
lemmas linorder_antisym_conv3 [where 'b = "?'a::linorder"] = linorder.antisym_conv3 |
|
342 |
lemmas leI [where 'b = "?'a::linorder"] = linorder.leI |
|
343 |
lemmas leD [where 'b = "?'a::linorder"] = linorder.leD |
|
344 |
lemmas not_leE [where 'b = "?'a::linorder"] = linorder.not_leE |
|
| 16796 | 345 |
|
| 21083 | 346 |
|
347 |
subsection {* Reasoning tools setup *}
|
|
348 |
||
| 21091 | 349 |
ML {*
|
350 |
local |
|
351 |
||
352 |
fun decomp_gen sort thy (Trueprop $ t) = |
|
| 21248 | 353 |
let |
354 |
fun of_sort t = |
|
355 |
let |
|
356 |
val T = type_of t |
|
357 |
in |
|
| 21091 | 358 |
(* exclude numeric types: linear arithmetic subsumes transitivity *) |
| 21248 | 359 |
T <> HOLogic.natT andalso T <> HOLogic.intT |
360 |
andalso T <> HOLogic.realT andalso Sign.of_sort thy (T, sort) |
|
361 |
end; |
|
362 |
fun dec (Const ("Not", _) $ t) = (case dec t
|
|
363 |
of NONE => NONE |
|
364 |
| SOME (t1, rel, t2) => SOME (t1, "~" ^ rel, t2)) |
|
365 |
| dec (Const ("op =", _) $ t1 $ t2) =
|
|
366 |
if of_sort t1 |
|
367 |
then SOME (t1, "=", t2) |
|
368 |
else NONE |
|
369 |
| dec (Const ("Orderings.less_eq", _) $ t1 $ t2) =
|
|
370 |
if of_sort t1 |
|
371 |
then SOME (t1, "<=", t2) |
|
372 |
else NONE |
|
373 |
| dec (Const ("Orderings.less", _) $ t1 $ t2) =
|
|
374 |
if of_sort t1 |
|
375 |
then SOME (t1, "<", t2) |
|
376 |
else NONE |
|
377 |
| dec _ = NONE; |
|
| 21091 | 378 |
in dec t end; |
379 |
||
380 |
in |
|
381 |
||
382 |
(* The setting up of Quasi_Tac serves as a demo. Since there is no |
|
383 |
class for quasi orders, the tactics Quasi_Tac.trans_tac and |
|
384 |
Quasi_Tac.quasi_tac are not of much use. *) |
|
385 |
||
| 21248 | 386 |
structure Quasi_Tac = Quasi_Tac_Fun ( |
387 |
struct |
|
388 |
val le_trans = thm "order_trans"; |
|
389 |
val le_refl = thm "order_refl"; |
|
390 |
val eqD1 = thm "order_eq_refl"; |
|
391 |
val eqD2 = thm "sym" RS thm "order_eq_refl"; |
|
392 |
val less_reflE = thm "order_less_irrefl" RS thm "notE"; |
|
393 |
val less_imp_le = thm "order_less_imp_le"; |
|
394 |
val le_neq_trans = thm "order_le_neq_trans"; |
|
395 |
val neq_le_trans = thm "order_neq_le_trans"; |
|
396 |
val less_imp_neq = thm "less_imp_neq"; |
|
397 |
val decomp_trans = decomp_gen ["Orderings.order"]; |
|
398 |
val decomp_quasi = decomp_gen ["Orderings.order"]; |
|
399 |
end); |
|
| 21091 | 400 |
|
401 |
structure Order_Tac = Order_Tac_Fun ( |
|
| 21248 | 402 |
struct |
403 |
val less_reflE = thm "order_less_irrefl" RS thm "notE"; |
|
404 |
val le_refl = thm "order_refl"; |
|
405 |
val less_imp_le = thm "order_less_imp_le"; |
|
406 |
val not_lessI = thm "linorder_not_less" RS thm "iffD2"; |
|
407 |
val not_leI = thm "linorder_not_le" RS thm "iffD2"; |
|
408 |
val not_lessD = thm "linorder_not_less" RS thm "iffD1"; |
|
409 |
val not_leD = thm "linorder_not_le" RS thm "iffD1"; |
|
410 |
val eqI = thm "order_antisym"; |
|
411 |
val eqD1 = thm "order_eq_refl"; |
|
412 |
val eqD2 = thm "sym" RS thm "order_eq_refl"; |
|
413 |
val less_trans = thm "order_less_trans"; |
|
414 |
val less_le_trans = thm "order_less_le_trans"; |
|
415 |
val le_less_trans = thm "order_le_less_trans"; |
|
416 |
val le_trans = thm "order_trans"; |
|
417 |
val le_neq_trans = thm "order_le_neq_trans"; |
|
418 |
val neq_le_trans = thm "order_neq_le_trans"; |
|
419 |
val less_imp_neq = thm "less_imp_neq"; |
|
420 |
val eq_neq_eq_imp_neq = thm "eq_neq_eq_imp_neq"; |
|
421 |
val not_sym = thm "not_sym"; |
|
422 |
val decomp_part = decomp_gen ["Orderings.order"]; |
|
423 |
val decomp_lin = decomp_gen ["Orderings.linorder"]; |
|
424 |
end); |
|
| 21091 | 425 |
|
426 |
end; |
|
427 |
*} |
|
428 |
||
| 21083 | 429 |
setup {*
|
430 |
let |
|
431 |
||
432 |
val order_antisym_conv = thm "order_antisym_conv" |
|
433 |
val linorder_antisym_conv1 = thm "linorder_antisym_conv1" |
|
434 |
val linorder_antisym_conv2 = thm "linorder_antisym_conv2" |
|
435 |
val linorder_antisym_conv3 = thm "linorder_antisym_conv3" |
|
436 |
||
437 |
fun prp t thm = (#prop (rep_thm thm) = t); |
|
| 15524 | 438 |
|
| 21083 | 439 |
fun prove_antisym_le sg ss ((le as Const(_,T)) $ r $ s) = |
440 |
let val prems = prems_of_ss ss; |
|
441 |
val less = Const("Orderings.less",T);
|
|
442 |
val t = HOLogic.mk_Trueprop(le $ s $ r); |
|
443 |
in case find_first (prp t) prems of |
|
444 |
NONE => |
|
445 |
let val t = HOLogic.mk_Trueprop(HOLogic.Not $ (less $ r $ s)) |
|
446 |
in case find_first (prp t) prems of |
|
447 |
NONE => NONE |
|
448 |
| SOME thm => SOME(mk_meta_eq(thm RS linorder_antisym_conv1)) |
|
449 |
end |
|
450 |
| SOME thm => SOME(mk_meta_eq(thm RS order_antisym_conv)) |
|
451 |
end |
|
452 |
handle THM _ => NONE; |
|
| 15524 | 453 |
|
| 21083 | 454 |
fun prove_antisym_less sg ss (NotC $ ((less as Const(_,T)) $ r $ s)) = |
455 |
let val prems = prems_of_ss ss; |
|
456 |
val le = Const("Orderings.less_eq",T);
|
|
457 |
val t = HOLogic.mk_Trueprop(le $ r $ s); |
|
458 |
in case find_first (prp t) prems of |
|
459 |
NONE => |
|
460 |
let val t = HOLogic.mk_Trueprop(NotC $ (less $ s $ r)) |
|
461 |
in case find_first (prp t) prems of |
|
462 |
NONE => NONE |
|
463 |
| SOME thm => SOME(mk_meta_eq(thm RS linorder_antisym_conv3)) |
|
464 |
end |
|
465 |
| SOME thm => SOME(mk_meta_eq(thm RS linorder_antisym_conv2)) |
|
466 |
end |
|
467 |
handle THM _ => NONE; |
|
| 15524 | 468 |
|
| 21248 | 469 |
fun add_simprocs procs thy = |
470 |
(Simplifier.change_simpset_of thy (fn ss => ss |
|
471 |
addsimprocs (map (fn (name, raw_ts, proc) => |
|
472 |
Simplifier.simproc thy name raw_ts proc)) procs); thy); |
|
473 |
fun add_solver name tac thy = |
|
474 |
(Simplifier.change_simpset_of thy (fn ss => ss addSolver |
|
475 |
(mk_solver name (K tac))); thy); |
|
| 21083 | 476 |
|
477 |
in |
|
| 21248 | 478 |
add_simprocs [ |
479 |
("antisym le", ["(x::'a::order) <= y"], prove_antisym_le),
|
|
480 |
("antisym less", ["~ (x::'a::linorder) < y"], prove_antisym_less)
|
|
481 |
] |
|
482 |
#> add_solver "Trans_linear" Order_Tac.linear_tac |
|
483 |
#> add_solver "Trans_partial" Order_Tac.partial_tac |
|
484 |
(* Adding the transitivity reasoners also as safe solvers showed a slight |
|
485 |
speed up, but the reasoning strength appears to be not higher (at least |
|
486 |
no breaking of additional proofs in the entire HOL distribution, as |
|
487 |
of 5 March 2004, was observed). *) |
|
| 21083 | 488 |
end |
489 |
*} |
|
| 15524 | 490 |
|
491 |
||
| 21083 | 492 |
subsection {* Bounded quantifiers *}
|
493 |
||
494 |
syntax |
|
|
21180
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
495 |
"_All_less" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
496 |
"_Ex_less" :: "[idt, 'a, bool] => bool" ("(3EX _<_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
497 |
"_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
498 |
"_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3EX _<=_./ _)" [0, 0, 10] 10)
|
| 21083 | 499 |
|
|
21180
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
500 |
"_All_greater" :: "[idt, 'a, bool] => bool" ("(3ALL _>_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
501 |
"_Ex_greater" :: "[idt, 'a, bool] => bool" ("(3EX _>_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
502 |
"_All_greater_eq" :: "[idt, 'a, bool] => bool" ("(3ALL _>=_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
503 |
"_Ex_greater_eq" :: "[idt, 'a, bool] => bool" ("(3EX _>=_./ _)" [0, 0, 10] 10)
|
| 21083 | 504 |
|
505 |
syntax (xsymbols) |
|
|
21180
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
506 |
"_All_less" :: "[idt, 'a, bool] => bool" ("(3\<forall>_<_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
507 |
"_Ex_less" :: "[idt, 'a, bool] => bool" ("(3\<exists>_<_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
508 |
"_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
509 |
"_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
|
| 21083 | 510 |
|
|
21180
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
511 |
"_All_greater" :: "[idt, 'a, bool] => bool" ("(3\<forall>_>_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
512 |
"_Ex_greater" :: "[idt, 'a, bool] => bool" ("(3\<exists>_>_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
513 |
"_All_greater_eq" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
514 |
"_Ex_greater_eq" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10)
|
| 21083 | 515 |
|
516 |
syntax (HOL) |
|
|
21180
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
517 |
"_All_less" :: "[idt, 'a, bool] => bool" ("(3! _<_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
518 |
"_Ex_less" :: "[idt, 'a, bool] => bool" ("(3? _<_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
519 |
"_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3! _<=_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
520 |
"_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3? _<=_./ _)" [0, 0, 10] 10)
|
| 21083 | 521 |
|
522 |
syntax (HTML output) |
|
|
21180
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
523 |
"_All_less" :: "[idt, 'a, bool] => bool" ("(3\<forall>_<_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
524 |
"_Ex_less" :: "[idt, 'a, bool] => bool" ("(3\<exists>_<_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
525 |
"_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
526 |
"_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
|
| 21083 | 527 |
|
|
21180
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
528 |
"_All_greater" :: "[idt, 'a, bool] => bool" ("(3\<forall>_>_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
529 |
"_Ex_greater" :: "[idt, 'a, bool] => bool" ("(3\<exists>_>_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
530 |
"_All_greater_eq" :: "[idt, 'a, bool] => bool" ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10)
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
531 |
"_Ex_greater_eq" :: "[idt, 'a, bool] => bool" ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10)
|
| 21083 | 532 |
|
533 |
translations |
|
534 |
"ALL x<y. P" => "ALL x. x < y \<longrightarrow> P" |
|
535 |
"EX x<y. P" => "EX x. x < y \<and> P" |
|
536 |
"ALL x<=y. P" => "ALL x. x <= y \<longrightarrow> P" |
|
537 |
"EX x<=y. P" => "EX x. x <= y \<and> P" |
|
538 |
"ALL x>y. P" => "ALL x. x > y \<longrightarrow> P" |
|
539 |
"EX x>y. P" => "EX x. x > y \<and> P" |
|
540 |
"ALL x>=y. P" => "ALL x. x >= y \<longrightarrow> P" |
|
541 |
"EX x>=y. P" => "EX x. x >= y \<and> P" |
|
542 |
||
543 |
print_translation {*
|
|
544 |
let |
|
|
21180
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
545 |
val syntax_name = Sign.const_syntax_name (the_context ()); |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
546 |
val impl = syntax_name "op -->"; |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
547 |
val conj = syntax_name "op &"; |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
548 |
val less = syntax_name "Orderings.less"; |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
549 |
val less_eq = syntax_name "Orderings.less_eq"; |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
550 |
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
551 |
val trans = |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
552 |
[(("ALL ", impl, less), ("_All_less", "_All_greater")),
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
553 |
(("ALL ", impl, less_eq), ("_All_less_eq", "_All_greater_eq")),
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
554 |
(("EX ", conj, less), ("_Ex_less", "_Ex_greater")),
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
555 |
(("EX ", conj, less_eq), ("_Ex_less_eq", "_Ex_greater_eq"))];
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
556 |
|
| 21083 | 557 |
fun mk v v' c n P = |
|
21180
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
558 |
if v = v' andalso not (Term.exists_subterm (fn Free (x, _) => x = v | _ => false) n) |
| 21083 | 559 |
then Syntax.const c $ Syntax.mark_bound v' $ n $ P else raise Match; |
|
21180
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
560 |
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
561 |
fun tr' q = (q, |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
562 |
fn [Const ("_bound", _) $ Free (v, _), Const (c, _) $ (Const (d, _) $ t $ u) $ P] =>
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
563 |
(case AList.lookup (op =) trans (q, c, d) of |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
564 |
NONE => raise Match |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
565 |
| SOME (l, g) => |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
566 |
(case (t, u) of |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
567 |
(Const ("_bound", _) $ Free (v', _), n) => mk v v' l n P
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
568 |
| (n, Const ("_bound", _) $ Free (v', _)) => mk v v' g n P
|
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
569 |
| _ => raise Match)) |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
570 |
| _ => raise Match); |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
571 |
in [tr' "ALL ", tr' "EX "] end |
| 21083 | 572 |
*} |
573 |
||
574 |
||
|
21383
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
575 |
subsection {* Transitivity reasoning *}
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
576 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
577 |
lemma ord_le_eq_trans: "a <= b ==> b = c ==> a <= c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
578 |
by (rule subst) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
579 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
580 |
lemma ord_eq_le_trans: "a = b ==> b <= c ==> a <= c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
581 |
by (rule ssubst) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
582 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
583 |
lemma ord_less_eq_trans: "a < b ==> b = c ==> a < c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
584 |
by (rule subst) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
585 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
586 |
lemma ord_eq_less_trans: "a = b ==> b < c ==> a < c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
587 |
by (rule ssubst) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
588 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
589 |
lemma order_less_subst2: "(a::'a::order) < b ==> f b < (c::'c::order) ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
590 |
(!!x y. x < y ==> f x < f y) ==> f a < c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
591 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
592 |
assume r: "!!x y. x < y ==> f x < f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
593 |
assume "a < b" hence "f a < f b" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
594 |
also assume "f b < c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
595 |
finally (order_less_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
596 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
597 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
598 |
lemma order_less_subst1: "(a::'a::order) < f b ==> (b::'b::order) < c ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
599 |
(!!x y. x < y ==> f x < f y) ==> a < f c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
600 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
601 |
assume r: "!!x y. x < y ==> f x < f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
602 |
assume "a < f b" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
603 |
also assume "b < c" hence "f b < f c" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
604 |
finally (order_less_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
605 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
606 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
607 |
lemma order_le_less_subst2: "(a::'a::order) <= b ==> f b < (c::'c::order) ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
608 |
(!!x y. x <= y ==> f x <= f y) ==> f a < c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
609 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
610 |
assume r: "!!x y. x <= y ==> f x <= f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
611 |
assume "a <= b" hence "f a <= f b" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
612 |
also assume "f b < c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
613 |
finally (order_le_less_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
614 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
615 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
616 |
lemma order_le_less_subst1: "(a::'a::order) <= f b ==> (b::'b::order) < c ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
617 |
(!!x y. x < y ==> f x < f y) ==> a < f c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
618 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
619 |
assume r: "!!x y. x < y ==> f x < f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
620 |
assume "a <= f b" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
621 |
also assume "b < c" hence "f b < f c" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
622 |
finally (order_le_less_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
623 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
624 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
625 |
lemma order_less_le_subst2: "(a::'a::order) < b ==> f b <= (c::'c::order) ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
626 |
(!!x y. x < y ==> f x < f y) ==> f a < c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
627 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
628 |
assume r: "!!x y. x < y ==> f x < f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
629 |
assume "a < b" hence "f a < f b" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
630 |
also assume "f b <= c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
631 |
finally (order_less_le_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
632 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
633 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
634 |
lemma order_less_le_subst1: "(a::'a::order) < f b ==> (b::'b::order) <= c ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
635 |
(!!x y. x <= y ==> f x <= f y) ==> a < f c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
636 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
637 |
assume r: "!!x y. x <= y ==> f x <= f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
638 |
assume "a < f b" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
639 |
also assume "b <= c" hence "f b <= f c" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
640 |
finally (order_less_le_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
641 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
642 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
643 |
lemma order_subst1: "(a::'a::order) <= f b ==> (b::'b::order) <= c ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
644 |
(!!x y. x <= y ==> f x <= f y) ==> a <= f c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
645 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
646 |
assume r: "!!x y. x <= y ==> f x <= f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
647 |
assume "a <= f b" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
648 |
also assume "b <= c" hence "f b <= f c" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
649 |
finally (order_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
650 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
651 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
652 |
lemma order_subst2: "(a::'a::order) <= b ==> f b <= (c::'c::order) ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
653 |
(!!x y. x <= y ==> f x <= f y) ==> f a <= c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
654 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
655 |
assume r: "!!x y. x <= y ==> f x <= f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
656 |
assume "a <= b" hence "f a <= f b" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
657 |
also assume "f b <= c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
658 |
finally (order_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
659 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
660 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
661 |
lemma ord_le_eq_subst: "a <= b ==> f b = c ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
662 |
(!!x y. x <= y ==> f x <= f y) ==> f a <= c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
663 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
664 |
assume r: "!!x y. x <= y ==> f x <= f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
665 |
assume "a <= b" hence "f a <= f b" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
666 |
also assume "f b = c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
667 |
finally (ord_le_eq_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
668 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
669 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
670 |
lemma ord_eq_le_subst: "a = f b ==> b <= c ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
671 |
(!!x y. x <= y ==> f x <= f y) ==> a <= f c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
672 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
673 |
assume r: "!!x y. x <= y ==> f x <= f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
674 |
assume "a = f b" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
675 |
also assume "b <= c" hence "f b <= f c" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
676 |
finally (ord_eq_le_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
677 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
678 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
679 |
lemma ord_less_eq_subst: "a < b ==> f b = c ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
680 |
(!!x y. x < y ==> f x < f y) ==> f a < c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
681 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
682 |
assume r: "!!x y. x < y ==> f x < f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
683 |
assume "a < b" hence "f a < f b" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
684 |
also assume "f b = c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
685 |
finally (ord_less_eq_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
686 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
687 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
688 |
lemma ord_eq_less_subst: "a = f b ==> b < c ==> |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
689 |
(!!x y. x < y ==> f x < f y) ==> a < f c" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
690 |
proof - |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
691 |
assume r: "!!x y. x < y ==> f x < f y" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
692 |
assume "a = f b" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
693 |
also assume "b < c" hence "f b < f c" by (rule r) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
694 |
finally (ord_eq_less_trans) show ?thesis . |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
695 |
qed |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
696 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
697 |
text {*
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
698 |
Note that this list of rules is in reverse order of priorities. |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
699 |
*} |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
700 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
701 |
lemmas order_trans_rules [trans] = |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
702 |
order_less_subst2 |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
703 |
order_less_subst1 |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
704 |
order_le_less_subst2 |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
705 |
order_le_less_subst1 |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
706 |
order_less_le_subst2 |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
707 |
order_less_le_subst1 |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
708 |
order_subst2 |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
709 |
order_subst1 |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
710 |
ord_le_eq_subst |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
711 |
ord_eq_le_subst |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
712 |
ord_less_eq_subst |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
713 |
ord_eq_less_subst |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
714 |
forw_subst |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
715 |
back_subst |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
716 |
rev_mp |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
717 |
mp |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
718 |
order_neq_le_trans |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
719 |
order_le_neq_trans |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
720 |
order_less_trans |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
721 |
order_less_asym' |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
722 |
order_le_less_trans |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
723 |
order_less_le_trans |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
724 |
order_trans |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
725 |
order_antisym |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
726 |
ord_le_eq_trans |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
727 |
ord_eq_le_trans |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
728 |
ord_less_eq_trans |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
729 |
ord_eq_less_trans |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
730 |
trans |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
731 |
|
| 21083 | 732 |
|
|
21180
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
733 |
(* FIXME cleanup *) |
|
f27f12bcafb8
fixed print_translation for ALL/EX and <, <=, etc.; tuned syntax names;
wenzelm
parents:
21091
diff
changeset
|
734 |
|
| 21083 | 735 |
text {* These support proving chains of decreasing inequalities
|
736 |
a >= b >= c ... in Isar proofs. *} |
|
737 |
||
738 |
lemma xt1: |
|
739 |
"a = b ==> b > c ==> a > c" |
|
740 |
"a > b ==> b = c ==> a > c" |
|
741 |
"a = b ==> b >= c ==> a >= c" |
|
742 |
"a >= b ==> b = c ==> a >= c" |
|
743 |
"(x::'a::order) >= y ==> y >= x ==> x = y" |
|
744 |
"(x::'a::order) >= y ==> y >= z ==> x >= z" |
|
745 |
"(x::'a::order) > y ==> y >= z ==> x > z" |
|
746 |
"(x::'a::order) >= y ==> y > z ==> x > z" |
|
747 |
"(a::'a::order) > b ==> b > a ==> ?P" |
|
748 |
"(x::'a::order) > y ==> y > z ==> x > z" |
|
749 |
"(a::'a::order) >= b ==> a ~= b ==> a > b" |
|
750 |
"(a::'a::order) ~= b ==> a >= b ==> a > b" |
|
751 |
"a = f b ==> b > c ==> (!!x y. x > y ==> f x > f y) ==> a > f c" |
|
752 |
"a > b ==> f b = c ==> (!!x y. x > y ==> f x > f y) ==> f a > c" |
|
753 |
"a = f b ==> b >= c ==> (!!x y. x >= y ==> f x >= f y) ==> a >= f c" |
|
754 |
"a >= b ==> f b = c ==> (!! x y. x >= y ==> f x >= f y) ==> f a >= c" |
|
755 |
by auto |
|
756 |
||
757 |
lemma xt2: |
|
758 |
"(a::'a::order) >= f b ==> b >= c ==> (!!x y. x >= y ==> f x >= f y) ==> a >= f c" |
|
759 |
by (subgoal_tac "f b >= f c", force, force) |
|
760 |
||
761 |
lemma xt3: "(a::'a::order) >= b ==> (f b::'b::order) >= c ==> |
|
762 |
(!!x y. x >= y ==> f x >= f y) ==> f a >= c" |
|
763 |
by (subgoal_tac "f a >= f b", force, force) |
|
764 |
||
765 |
lemma xt4: "(a::'a::order) > f b ==> (b::'b::order) >= c ==> |
|
766 |
(!!x y. x >= y ==> f x >= f y) ==> a > f c" |
|
767 |
by (subgoal_tac "f b >= f c", force, force) |
|
768 |
||
769 |
lemma xt5: "(a::'a::order) > b ==> (f b::'b::order) >= c==> |
|
770 |
(!!x y. x > y ==> f x > f y) ==> f a > c" |
|
771 |
by (subgoal_tac "f a > f b", force, force) |
|
772 |
||
773 |
lemma xt6: "(a::'a::order) >= f b ==> b > c ==> |
|
774 |
(!!x y. x > y ==> f x > f y) ==> a > f c" |
|
775 |
by (subgoal_tac "f b > f c", force, force) |
|
776 |
||
777 |
lemma xt7: "(a::'a::order) >= b ==> (f b::'b::order) > c ==> |
|
778 |
(!!x y. x >= y ==> f x >= f y) ==> f a > c" |
|
779 |
by (subgoal_tac "f a >= f b", force, force) |
|
780 |
||
781 |
lemma xt8: "(a::'a::order) > f b ==> (b::'b::order) > c ==> |
|
782 |
(!!x y. x > y ==> f x > f y) ==> a > f c" |
|
783 |
by (subgoal_tac "f b > f c", force, force) |
|
784 |
||
785 |
lemma xt9: "(a::'a::order) > b ==> (f b::'b::order) > c ==> |
|
786 |
(!!x y. x > y ==> f x > f y) ==> f a > c" |
|
787 |
by (subgoal_tac "f a > f b", force, force) |
|
788 |
||
789 |
lemmas xtrans = xt1 xt2 xt3 xt4 xt5 xt6 xt7 xt8 xt9 |
|
790 |
||
791 |
(* |
|
792 |
Since "a >= b" abbreviates "b <= a", the abbreviation "..." stands |
|
793 |
for the wrong thing in an Isar proof. |
|
794 |
||
795 |
The extra transitivity rules can be used as follows: |
|
796 |
||
797 |
lemma "(a::'a::order) > z" |
|
798 |
proof - |
|
799 |
have "a >= b" (is "_ >= ?rhs") |
|
800 |
sorry |
|
801 |
also have "?rhs >= c" (is "_ >= ?rhs") |
|
802 |
sorry |
|
803 |
also (xtrans) have "?rhs = d" (is "_ = ?rhs") |
|
804 |
sorry |
|
805 |
also (xtrans) have "?rhs >= e" (is "_ >= ?rhs") |
|
806 |
sorry |
|
807 |
also (xtrans) have "?rhs > f" (is "_ > ?rhs") |
|
808 |
sorry |
|
809 |
also (xtrans) have "?rhs > z" |
|
810 |
sorry |
|
811 |
finally (xtrans) show ?thesis . |
|
812 |
qed |
|
813 |
||
814 |
Alternatively, one can use "declare xtrans [trans]" and then |
|
815 |
leave out the "(xtrans)" above. |
|
816 |
*) |
|
817 |
||
|
21383
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
818 |
subsection {* Monotonicity, syntactic least value operator and min/max *}
|
| 21083 | 819 |
|
|
21216
1c8580913738
made locale partial_order compatible with axclass order; changed import order; consecutive changes
haftmann
parents:
21204
diff
changeset
|
820 |
locale mono = |
|
1c8580913738
made locale partial_order compatible with axclass order; changed import order; consecutive changes
haftmann
parents:
21204
diff
changeset
|
821 |
fixes f |
|
1c8580913738
made locale partial_order compatible with axclass order; changed import order; consecutive changes
haftmann
parents:
21204
diff
changeset
|
822 |
assumes mono: "A \<le> B \<Longrightarrow> f A \<le> f B" |
|
1c8580913738
made locale partial_order compatible with axclass order; changed import order; consecutive changes
haftmann
parents:
21204
diff
changeset
|
823 |
|
|
1c8580913738
made locale partial_order compatible with axclass order; changed import order; consecutive changes
haftmann
parents:
21204
diff
changeset
|
824 |
lemmas monoI [intro?] = mono.intro |
|
1c8580913738
made locale partial_order compatible with axclass order; changed import order; consecutive changes
haftmann
parents:
21204
diff
changeset
|
825 |
and monoD [dest?] = mono.mono |
| 21083 | 826 |
|
827 |
constdefs |
|
828 |
Least :: "('a::ord => bool) => 'a" (binder "LEAST " 10)
|
|
829 |
"Least P == THE x. P x & (ALL y. P y --> x <= y)" |
|
830 |
-- {* We can no longer use LeastM because the latter requires Hilbert-AC. *}
|
|
831 |
||
|
21383
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
832 |
lemma LeastI2_order: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
833 |
"[| P (x::'a::order); |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
834 |
!!y. P y ==> x <= y; |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
835 |
!!x. [| P x; ALL y. P y --> x \<le> y |] ==> Q x |] |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
836 |
==> Q (Least P)" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
837 |
apply (unfold Least_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
838 |
apply (rule theI2) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
839 |
apply (blast intro: order_antisym)+ |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
840 |
done |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
841 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
842 |
lemma Least_equality: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
843 |
"[| P (k::'a::order); !!x. P x ==> k <= x |] ==> (LEAST x. P x) = k" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
844 |
apply (simp add: Least_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
845 |
apply (rule the_equality) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
846 |
apply (auto intro!: order_antisym) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
847 |
done |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
848 |
|
| 21083 | 849 |
constdefs |
850 |
min :: "['a::ord, 'a] => 'a" |
|
851 |
"min a b == (if a <= b then a else b)" |
|
852 |
max :: "['a::ord, 'a] => 'a" |
|
853 |
"max a b == (if a <= b then b else a)" |
|
854 |
||
|
21383
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
855 |
lemma min_linorder: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
856 |
"linorder.min (op \<le> \<Colon> 'a\<Colon>linorder \<Rightarrow> 'a \<Rightarrow> bool) = min" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
857 |
by (rule+) (simp add: min_def linorder.min_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
858 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
859 |
lemma max_linorder: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
860 |
"linorder.max (op \<le> \<Colon> 'a\<Colon>linorder \<Rightarrow> 'a \<Rightarrow> bool) = max" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
861 |
by (rule+) (simp add: max_def linorder.max_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
862 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
863 |
lemmas min_le_iff_disj = linorder.min_le_iff_disj [where 'b = "?'a::linorder", simplified min_linorder] |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
864 |
lemmas le_max_iff_disj = linorder.le_max_iff_disj [where 'b = "?'a::linorder", simplified max_linorder] |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
865 |
lemmas min_less_iff_disj = linorder.min_less_iff_disj [where 'b = "?'a::linorder", simplified min_linorder] |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
866 |
lemmas less_max_iff_disj = linorder.less_max_iff_disj [where 'b = "?'a::linorder", simplified max_linorder] |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
867 |
lemmas min_less_iff_conj [simp] = linorder.min_less_iff_conj [where 'b = "?'a::linorder", simplified min_linorder] |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
868 |
lemmas max_less_iff_conj [simp] = linorder.max_less_iff_conj [where 'b = "?'a::linorder", simplified max_linorder] |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
869 |
lemmas split_min = linorder.split_min [where 'b = "?'a::linorder", simplified min_linorder] |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
870 |
lemmas split_max = linorder.split_max [where 'b = "?'a::linorder", simplified max_linorder] |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
871 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
872 |
lemma min_leastL: "(!!x. least <= x) ==> min least x = least" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
873 |
by (simp add: min_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
874 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
875 |
lemma max_leastL: "(!!x. least <= x) ==> max least x = x" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
876 |
by (simp add: max_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
877 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
878 |
lemma min_leastR: "(\<And>x\<Colon>'a\<Colon>order. least \<le> x) \<Longrightarrow> min x least = least" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
879 |
apply (simp add: min_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
880 |
apply (blast intro: order_antisym) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
881 |
done |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
882 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
883 |
lemma max_leastR: "(\<And>x\<Colon>'a\<Colon>order. least \<le> x) \<Longrightarrow> max x least = x" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
884 |
apply (simp add: max_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
885 |
apply (blast intro: order_antisym) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
886 |
done |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
887 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
888 |
lemma min_of_mono: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
889 |
"(!!x y. (f x <= f y) = (x <= y)) ==> min (f m) (f n) = f (min m n)" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
890 |
by (simp add: min_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
891 |
|
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
892 |
lemma max_of_mono: |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
893 |
"(!!x y. (f x <= f y) = (x <= y)) ==> max (f m) (f n) = f (max m n)" |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
894 |
by (simp add: max_def) |
|
17e6275e13f5
added transitivity rules, reworking of min/max lemmas
haftmann
parents:
21329
diff
changeset
|
895 |
|
| 15524 | 896 |
end |