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