author | wenzelm |
Sun, 22 Dec 2019 15:48:42 +0100 | |
changeset 71333 | c898cd5b8519 |
parent 70184 | a7aba6db79a1 |
child 72384 | b037517c815b |
permissions | -rw-r--r-- |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
1 |
(* Title: HOL/Lattices_Big.thy |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
2 |
Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
3 |
with contributions by Jeremy Avigad |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
4 |
*) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
5 |
|
60758 | 6 |
section \<open>Big infimum (minimum) and supremum (maximum) over finite (non-empty) sets\<close> |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
7 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
8 |
theory Lattices_Big |
66364 | 9 |
imports Option |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
10 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
11 |
|
60758 | 12 |
subsection \<open>Generic lattice operations over a set\<close> |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
13 |
|
60758 | 14 |
subsubsection \<open>Without neutral element\<close> |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
15 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
16 |
locale semilattice_set = semilattice |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
17 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
18 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
19 |
interpretation comp_fun_idem f |
61169 | 20 |
by standard (simp_all add: fun_eq_iff left_commute) |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
21 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
22 |
definition F :: "'a set \<Rightarrow> 'a" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
23 |
where |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
24 |
eq_fold': "F A = the (Finite_Set.fold (\<lambda>x y. Some (case y of None \<Rightarrow> x | Some z \<Rightarrow> f x z)) None A)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
25 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
26 |
lemma eq_fold: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
27 |
assumes "finite A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
28 |
shows "F (insert x A) = Finite_Set.fold f x A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
29 |
proof (rule sym) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
30 |
let ?f = "\<lambda>x y. Some (case y of None \<Rightarrow> x | Some z \<Rightarrow> f x z)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
31 |
interpret comp_fun_idem "?f" |
61169 | 32 |
by standard (simp_all add: fun_eq_iff commute left_commute split: option.split) |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
33 |
from assms show "Finite_Set.fold f x A = F (insert x A)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
34 |
proof induct |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
35 |
case empty then show ?case by (simp add: eq_fold') |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
36 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
37 |
case (insert y B) then show ?case by (simp add: insert_commute [of x] eq_fold') |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
38 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
39 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
40 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
41 |
lemma singleton [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
42 |
"F {x} = x" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
43 |
by (simp add: eq_fold) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
44 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
45 |
lemma insert_not_elem: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
46 |
assumes "finite A" and "x \<notin> A" and "A \<noteq> {}" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
47 |
shows "F (insert x A) = x \<^bold>* F A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
48 |
proof - |
60758 | 49 |
from \<open>A \<noteq> {}\<close> obtain b where "b \<in> A" by blast |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
50 |
then obtain B where *: "A = insert b B" "b \<notin> B" by (blast dest: mk_disjoint_insert) |
60758 | 51 |
with \<open>finite A\<close> and \<open>x \<notin> A\<close> |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
52 |
have "finite (insert x B)" and "b \<notin> insert x B" by auto |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
53 |
then have "F (insert b (insert x B)) = x \<^bold>* F (insert b B)" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
54 |
by (simp add: eq_fold) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
55 |
then show ?thesis by (simp add: * insert_commute) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
56 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
57 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
58 |
lemma in_idem: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
59 |
assumes "finite A" and "x \<in> A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
60 |
shows "x \<^bold>* F A = F A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
61 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
62 |
from assms have "A \<noteq> {}" by auto |
60758 | 63 |
with \<open>finite A\<close> show ?thesis using \<open>x \<in> A\<close> |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
64 |
by (induct A rule: finite_ne_induct) (auto simp add: ac_simps insert_not_elem) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
65 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
66 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
67 |
lemma insert [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
68 |
assumes "finite A" and "A \<noteq> {}" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
69 |
shows "F (insert x A) = x \<^bold>* F A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
70 |
using assms by (cases "x \<in> A") (simp_all add: insert_absorb in_idem insert_not_elem) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
71 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
72 |
lemma union: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
73 |
assumes "finite A" "A \<noteq> {}" and "finite B" "B \<noteq> {}" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
74 |
shows "F (A \<union> B) = F A \<^bold>* F B" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
75 |
using assms by (induct A rule: finite_ne_induct) (simp_all add: ac_simps) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
76 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
77 |
lemma remove: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
78 |
assumes "finite A" and "x \<in> A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
79 |
shows "F A = (if A - {x} = {} then x else x \<^bold>* F (A - {x}))" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
80 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
81 |
from assms obtain B where "A = insert x B" and "x \<notin> B" by (blast dest: mk_disjoint_insert) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
82 |
with assms show ?thesis by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
83 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
84 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
85 |
lemma insert_remove: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
86 |
assumes "finite A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
87 |
shows "F (insert x A) = (if A - {x} = {} then x else x \<^bold>* F (A - {x}))" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
88 |
using assms by (cases "x \<in> A") (simp_all add: insert_absorb remove) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
89 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
90 |
lemma subset: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
91 |
assumes "finite A" "B \<noteq> {}" and "B \<subseteq> A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
92 |
shows "F B \<^bold>* F A = F A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
93 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
94 |
from assms have "A \<noteq> {}" and "finite B" by (auto dest: finite_subset) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
95 |
with assms show ?thesis by (simp add: union [symmetric] Un_absorb1) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
96 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
97 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
98 |
lemma closed: |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
99 |
assumes "finite A" "A \<noteq> {}" and elem: "\<And>x y. x \<^bold>* y \<in> {x, y}" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
100 |
shows "F A \<in> A" |
60758 | 101 |
using \<open>finite A\<close> \<open>A \<noteq> {}\<close> proof (induct rule: finite_ne_induct) |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
102 |
case singleton then show ?case by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
103 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
104 |
case insert with elem show ?case by force |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
105 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
106 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
107 |
lemma hom_commute: |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
108 |
assumes hom: "\<And>x y. h (x \<^bold>* y) = h x \<^bold>* h y" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
109 |
and N: "finite N" "N \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
110 |
shows "h (F N) = F (h ` N)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
111 |
using N proof (induct rule: finite_ne_induct) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
112 |
case singleton thus ?case by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
113 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
114 |
case (insert n N) |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
115 |
then have "h (F (insert n N)) = h (n \<^bold>* F N)" by simp |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
116 |
also have "\<dots> = h n \<^bold>* h (F N)" by (rule hom) |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
117 |
also have "h (F N) = F (h ` N)" by (rule insert) |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
118 |
also have "h n \<^bold>* \<dots> = F (insert (h n) (h ` N))" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
119 |
using insert by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
120 |
also have "insert (h n) (h ` N) = h ` insert n N" by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
121 |
finally show ?case . |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
122 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
123 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56140
diff
changeset
|
124 |
lemma infinite: "\<not> finite A \<Longrightarrow> F A = the None" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56140
diff
changeset
|
125 |
unfolding eq_fold' by (cases "finite (UNIV::'a set)") (auto intro: finite_subset fold_infinite) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56140
diff
changeset
|
126 |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
127 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
128 |
|
54745 | 129 |
locale semilattice_order_set = binary?: semilattice_order + semilattice_set |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
130 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
131 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
132 |
lemma bounded_iff: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
133 |
assumes "finite A" and "A \<noteq> {}" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
134 |
shows "x \<^bold>\<le> F A \<longleftrightarrow> (\<forall>a\<in>A. x \<^bold>\<le> a)" |
67525
5d04d7bcd5f6
avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents:
67169
diff
changeset
|
135 |
using assms by (induct rule: finite_ne_induct) simp_all |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
136 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
137 |
lemma boundedI: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
138 |
assumes "finite A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
139 |
assumes "A \<noteq> {}" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
140 |
assumes "\<And>a. a \<in> A \<Longrightarrow> x \<^bold>\<le> a" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
141 |
shows "x \<^bold>\<le> F A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
142 |
using assms by (simp add: bounded_iff) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
143 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
144 |
lemma boundedE: |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
145 |
assumes "finite A" and "A \<noteq> {}" and "x \<^bold>\<le> F A" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
146 |
obtains "\<And>a. a \<in> A \<Longrightarrow> x \<^bold>\<le> a" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
147 |
using assms by (simp add: bounded_iff) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
148 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
149 |
lemma coboundedI: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
150 |
assumes "finite A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
151 |
and "a \<in> A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
152 |
shows "F A \<^bold>\<le> a" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
153 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
154 |
from assms have "A \<noteq> {}" by auto |
60758 | 155 |
from \<open>finite A\<close> \<open>A \<noteq> {}\<close> \<open>a \<in> A\<close> show ?thesis |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
156 |
proof (induct rule: finite_ne_induct) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
157 |
case singleton thus ?case by (simp add: refl) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
158 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
159 |
case (insert x B) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
160 |
from insert have "a = x \<or> a \<in> B" by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
161 |
then show ?case using insert by (auto intro: coboundedI2) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
162 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
163 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
164 |
|
67525
5d04d7bcd5f6
avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents:
67169
diff
changeset
|
165 |
lemma subset_imp: |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
166 |
assumes "A \<subseteq> B" and "A \<noteq> {}" and "finite B" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
167 |
shows "F B \<^bold>\<le> F A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
168 |
proof (cases "A = B") |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
169 |
case True then show ?thesis by (simp add: refl) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
170 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
171 |
case False |
60758 | 172 |
have B: "B = A \<union> (B - A)" using \<open>A \<subseteq> B\<close> by blast |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
173 |
then have "F B = F (A \<union> (B - A))" by simp |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
174 |
also have "\<dots> = F A \<^bold>* F (B - A)" using False assms by (subst union) (auto intro: finite_subset) |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
175 |
also have "\<dots> \<^bold>\<le> F A" by simp |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
176 |
finally show ?thesis . |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
177 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
178 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
179 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
180 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
181 |
|
60758 | 182 |
subsubsection \<open>With neutral element\<close> |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
183 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
184 |
locale semilattice_neutr_set = semilattice_neutr |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
185 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
186 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
187 |
interpretation comp_fun_idem f |
61169 | 188 |
by standard (simp_all add: fun_eq_iff left_commute) |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
189 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
190 |
definition F :: "'a set \<Rightarrow> 'a" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
191 |
where |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
192 |
eq_fold: "F A = Finite_Set.fold f \<^bold>1 A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
193 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
194 |
lemma infinite [simp]: |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
195 |
"\<not> finite A \<Longrightarrow> F A = \<^bold>1" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
196 |
by (simp add: eq_fold) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
197 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
198 |
lemma empty [simp]: |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
199 |
"F {} = \<^bold>1" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
200 |
by (simp add: eq_fold) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
201 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
202 |
lemma insert [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
203 |
assumes "finite A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
204 |
shows "F (insert x A) = x \<^bold>* F A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
205 |
using assms by (simp add: eq_fold) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
206 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
207 |
lemma in_idem: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
208 |
assumes "finite A" and "x \<in> A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
209 |
shows "x \<^bold>* F A = F A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
210 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
211 |
from assms have "A \<noteq> {}" by auto |
60758 | 212 |
with \<open>finite A\<close> show ?thesis using \<open>x \<in> A\<close> |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
213 |
by (induct A rule: finite_ne_induct) (auto simp add: ac_simps) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
214 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
215 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
216 |
lemma union: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
217 |
assumes "finite A" and "finite B" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
218 |
shows "F (A \<union> B) = F A \<^bold>* F B" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
219 |
using assms by (induct A) (simp_all add: ac_simps) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
220 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
221 |
lemma remove: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
222 |
assumes "finite A" and "x \<in> A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
223 |
shows "F A = x \<^bold>* F (A - {x})" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
224 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
225 |
from assms obtain B where "A = insert x B" and "x \<notin> B" by (blast dest: mk_disjoint_insert) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
226 |
with assms show ?thesis by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
227 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
228 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
229 |
lemma insert_remove: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
230 |
assumes "finite A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
231 |
shows "F (insert x A) = x \<^bold>* F (A - {x})" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
232 |
using assms by (cases "x \<in> A") (simp_all add: insert_absorb remove) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
233 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
234 |
lemma subset: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
235 |
assumes "finite A" and "B \<subseteq> A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
236 |
shows "F B \<^bold>* F A = F A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
237 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
238 |
from assms have "finite B" by (auto dest: finite_subset) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
239 |
with assms show ?thesis by (simp add: union [symmetric] Un_absorb1) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
240 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
241 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
242 |
lemma closed: |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
243 |
assumes "finite A" "A \<noteq> {}" and elem: "\<And>x y. x \<^bold>* y \<in> {x, y}" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
244 |
shows "F A \<in> A" |
60758 | 245 |
using \<open>finite A\<close> \<open>A \<noteq> {}\<close> proof (induct rule: finite_ne_induct) |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
246 |
case singleton then show ?case by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
247 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
248 |
case insert with elem show ?case by force |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
249 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
250 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
251 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
252 |
|
54745 | 253 |
locale semilattice_order_neutr_set = binary?: semilattice_neutr_order + semilattice_neutr_set |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
254 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
255 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
256 |
lemma bounded_iff: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
257 |
assumes "finite A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
258 |
shows "x \<^bold>\<le> F A \<longleftrightarrow> (\<forall>a\<in>A. x \<^bold>\<le> a)" |
67525
5d04d7bcd5f6
avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents:
67169
diff
changeset
|
259 |
using assms by (induct A) simp_all |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
260 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
261 |
lemma boundedI: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
262 |
assumes "finite A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
263 |
assumes "\<And>a. a \<in> A \<Longrightarrow> x \<^bold>\<le> a" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
264 |
shows "x \<^bold>\<le> F A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
265 |
using assms by (simp add: bounded_iff) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
266 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
267 |
lemma boundedE: |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
268 |
assumes "finite A" and "x \<^bold>\<le> F A" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
269 |
obtains "\<And>a. a \<in> A \<Longrightarrow> x \<^bold>\<le> a" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
270 |
using assms by (simp add: bounded_iff) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
271 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
272 |
lemma coboundedI: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
273 |
assumes "finite A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
274 |
and "a \<in> A" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
275 |
shows "F A \<^bold>\<le> a" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
276 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
277 |
from assms have "A \<noteq> {}" by auto |
60758 | 278 |
from \<open>finite A\<close> \<open>A \<noteq> {}\<close> \<open>a \<in> A\<close> show ?thesis |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
279 |
proof (induct rule: finite_ne_induct) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
280 |
case singleton thus ?case by (simp add: refl) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
281 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
282 |
case (insert x B) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
283 |
from insert have "a = x \<or> a \<in> B" by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
284 |
then show ?case using insert by (auto intro: coboundedI2) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
285 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
286 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
287 |
|
67525
5d04d7bcd5f6
avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents:
67169
diff
changeset
|
288 |
lemma subset_imp: |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
289 |
assumes "A \<subseteq> B" and "finite B" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
290 |
shows "F B \<^bold>\<le> F A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
291 |
proof (cases "A = B") |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
292 |
case True then show ?thesis by (simp add: refl) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
293 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
294 |
case False |
60758 | 295 |
have B: "B = A \<union> (B - A)" using \<open>A \<subseteq> B\<close> by blast |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
296 |
then have "F B = F (A \<union> (B - A))" by simp |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
297 |
also have "\<dots> = F A \<^bold>* F (B - A)" using False assms by (subst union) (auto intro: finite_subset) |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61776
diff
changeset
|
298 |
also have "\<dots> \<^bold>\<le> F A" by simp |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
299 |
finally show ?thesis . |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
300 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
301 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
302 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
303 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
304 |
|
60758 | 305 |
subsection \<open>Lattice operations on finite sets\<close> |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
306 |
|
54868 | 307 |
context semilattice_inf |
308 |
begin |
|
309 |
||
61605 | 310 |
sublocale Inf_fin: semilattice_order_set inf less_eq less |
61776 | 311 |
defines |
68980
5717fbc55521
added spaces because otherwise nonatomic arguments look awful: BIGf x -> BIG f x
nipkow
parents:
68801
diff
changeset
|
312 |
Inf_fin ("\<Sqinter>\<^sub>f\<^sub>i\<^sub>n _" [900] 900) = Inf_fin.F .. |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
313 |
|
54868 | 314 |
end |
315 |
||
316 |
context semilattice_sup |
|
317 |
begin |
|
318 |
||
61605 | 319 |
sublocale Sup_fin: semilattice_order_set sup greater_eq greater |
61776 | 320 |
defines |
68980
5717fbc55521
added spaces because otherwise nonatomic arguments look awful: BIGf x -> BIG f x
nipkow
parents:
68801
diff
changeset
|
321 |
Sup_fin ("\<Squnion>\<^sub>f\<^sub>i\<^sub>n _" [900] 900) = Sup_fin.F .. |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
322 |
|
54868 | 323 |
end |
324 |
||
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
325 |
|
60758 | 326 |
subsection \<open>Infimum and Supremum over non-empty sets\<close> |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
327 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
328 |
context lattice |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
329 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
330 |
|
54745 | 331 |
lemma Inf_fin_le_Sup_fin [simp]: |
332 |
assumes "finite A" and "A \<noteq> {}" |
|
333 |
shows "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<le> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA" |
|
334 |
proof - |
|
60758 | 335 |
from \<open>A \<noteq> {}\<close> obtain a where "a \<in> A" by blast |
336 |
with \<open>finite A\<close> have "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<le> a" by (rule Inf_fin.coboundedI) |
|
337 |
moreover from \<open>finite A\<close> \<open>a \<in> A\<close> have "a \<le> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA" by (rule Sup_fin.coboundedI) |
|
54745 | 338 |
ultimately show ?thesis by (rule order_trans) |
339 |
qed |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
340 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
341 |
lemma sup_Inf_absorb [simp]: |
54745 | 342 |
"finite A \<Longrightarrow> a \<in> A \<Longrightarrow> \<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<squnion> a = a" |
343 |
by (rule sup_absorb2) (rule Inf_fin.coboundedI) |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
344 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
345 |
lemma inf_Sup_absorb [simp]: |
54745 | 346 |
"finite A \<Longrightarrow> a \<in> A \<Longrightarrow> a \<sqinter> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA = a" |
347 |
by (rule inf_absorb1) (rule Sup_fin.coboundedI) |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
348 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
349 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
350 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
351 |
context distrib_lattice |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
352 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
353 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
354 |
lemma sup_Inf1_distrib: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
355 |
assumes "finite A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
356 |
and "A \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
357 |
shows "sup x (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA) = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup x a|a. a \<in> A}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
358 |
using assms by (simp add: image_def Inf_fin.hom_commute [where h="sup x", OF sup_inf_distrib1]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
359 |
(rule arg_cong [where f="Inf_fin"], blast) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
360 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
361 |
lemma sup_Inf2_distrib: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
362 |
assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
363 |
shows "sup (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB) = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup a b|a b. a \<in> A \<and> b \<in> B}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
364 |
using A proof (induct rule: finite_ne_induct) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
365 |
case singleton then show ?case |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
366 |
by (simp add: sup_Inf1_distrib [OF B]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
367 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
368 |
case (insert x A) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
369 |
have finB: "finite {sup x b |b. b \<in> B}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
370 |
by (rule finite_surj [where f = "sup x", OF B(1)], auto) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
371 |
have finAB: "finite {sup a b |a b. a \<in> A \<and> b \<in> B}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
372 |
proof - |
69276 | 373 |
have "{sup a b |a b. a \<in> A \<and> b \<in> B} = (\<Union>a\<in>A. \<Union>b\<in>B. {sup a b})" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
374 |
by blast |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
375 |
thus ?thesis by(simp add: insert(1) B(1)) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
376 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
377 |
have ne: "{sup a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
378 |
have "sup (\<Sqinter>\<^sub>f\<^sub>i\<^sub>n(insert x A)) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB) = sup (inf x (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA)) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
379 |
using insert by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
380 |
also have "\<dots> = inf (sup x (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB)) (sup (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB))" by(rule sup_inf_distrib2) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
381 |
also have "\<dots> = inf (\<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup x b|b. b \<in> B}) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup a b|a b. a \<in> A \<and> b \<in> B})" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
382 |
using insert by(simp add:sup_Inf1_distrib[OF B]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
383 |
also have "\<dots> = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n({sup x b |b. b \<in> B} \<union> {sup a b |a b. a \<in> A \<and> b \<in> B})" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
384 |
(is "_ = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n?M") |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
385 |
using B insert |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
386 |
by (simp add: Inf_fin.union [OF finB _ finAB ne]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
387 |
also have "?M = {sup a b |a b. a \<in> insert x A \<and> b \<in> B}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
388 |
by blast |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
389 |
finally show ?case . |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
390 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
391 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
392 |
lemma inf_Sup1_distrib: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
393 |
assumes "finite A" and "A \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
394 |
shows "inf x (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA) = \<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf x a|a. a \<in> A}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
395 |
using assms by (simp add: image_def Sup_fin.hom_commute [where h="inf x", OF inf_sup_distrib1]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
396 |
(rule arg_cong [where f="Sup_fin"], blast) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
397 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
398 |
lemma inf_Sup2_distrib: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
399 |
assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
400 |
shows "inf (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB) = \<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf a b|a b. a \<in> A \<and> b \<in> B}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
401 |
using A proof (induct rule: finite_ne_induct) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
402 |
case singleton thus ?case |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
403 |
by(simp add: inf_Sup1_distrib [OF B]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
404 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
405 |
case (insert x A) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
406 |
have finB: "finite {inf x b |b. b \<in> B}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
407 |
by(rule finite_surj[where f = "%b. inf x b", OF B(1)], auto) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
408 |
have finAB: "finite {inf a b |a b. a \<in> A \<and> b \<in> B}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
409 |
proof - |
69276 | 410 |
have "{inf a b |a b. a \<in> A \<and> b \<in> B} = (\<Union>a\<in>A. \<Union>b\<in>B. {inf a b})" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
411 |
by blast |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
412 |
thus ?thesis by(simp add: insert(1) B(1)) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
413 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
414 |
have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
415 |
have "inf (\<Squnion>\<^sub>f\<^sub>i\<^sub>n(insert x A)) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB) = inf (sup x (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA)) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
416 |
using insert by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
417 |
also have "\<dots> = sup (inf x (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB)) (inf (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB))" by(rule inf_sup_distrib2) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
418 |
also have "\<dots> = sup (\<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf x b|b. b \<in> B}) (\<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf a b|a b. a \<in> A \<and> b \<in> B})" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
419 |
using insert by(simp add:inf_Sup1_distrib[OF B]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
420 |
also have "\<dots> = \<Squnion>\<^sub>f\<^sub>i\<^sub>n({inf x b |b. b \<in> B} \<union> {inf a b |a b. a \<in> A \<and> b \<in> B})" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
421 |
(is "_ = \<Squnion>\<^sub>f\<^sub>i\<^sub>n?M") |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
422 |
using B insert |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
423 |
by (simp add: Sup_fin.union [OF finB _ finAB ne]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
424 |
also have "?M = {inf a b |a b. a \<in> insert x A \<and> b \<in> B}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
425 |
by blast |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
426 |
finally show ?case . |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
427 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
428 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
429 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
430 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
431 |
context complete_lattice |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
432 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
433 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
434 |
lemma Inf_fin_Inf: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
435 |
assumes "finite A" and "A \<noteq> {}" |
54868 | 436 |
shows "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA = \<Sqinter>A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
437 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
438 |
from assms obtain b B where "A = insert b B" and "finite B" by auto |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
439 |
then show ?thesis |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
440 |
by (simp add: Inf_fin.eq_fold inf_Inf_fold_inf inf.commute [of b]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
441 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
442 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
443 |
lemma Sup_fin_Sup: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
444 |
assumes "finite A" and "A \<noteq> {}" |
54868 | 445 |
shows "\<Squnion>\<^sub>f\<^sub>i\<^sub>nA = \<Squnion>A" |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
446 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
447 |
from assms obtain b B where "A = insert b B" and "finite B" by auto |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
448 |
then show ?thesis |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
449 |
by (simp add: Sup_fin.eq_fold sup_Sup_fold_sup sup.commute [of b]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
450 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
451 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
452 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
453 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
454 |
|
60758 | 455 |
subsection \<open>Minimum and Maximum over non-empty sets\<close> |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
456 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
457 |
context linorder |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
458 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
459 |
|
61605 | 460 |
sublocale Min: semilattice_order_set min less_eq less |
461 |
+ Max: semilattice_order_set max greater_eq greater |
|
61776 | 462 |
defines |
463 |
Min = Min.F and Max = Max.F .. |
|
54864
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
464 |
|
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
465 |
end |
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
466 |
|
67969
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
467 |
syntax |
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
468 |
"_MIN1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3MIN _./ _)" [0, 10] 10) |
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
469 |
"_MIN" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3MIN _\<in>_./ _)" [0, 0, 10] 10) |
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
470 |
"_MAX1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3MAX _./ _)" [0, 10] 10) |
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
471 |
"_MAX" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3MAX _\<in>_./ _)" [0, 0, 10] 10) |
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
472 |
|
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
473 |
translations |
68796
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents:
68795
diff
changeset
|
474 |
"MIN x y. f" \<rightleftharpoons> "MIN x. MIN y. f" |
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents:
68795
diff
changeset
|
475 |
"MIN x. f" \<rightleftharpoons> "CONST Min (CONST range (\<lambda>x. f))" |
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents:
68795
diff
changeset
|
476 |
"MIN x\<in>A. f" \<rightleftharpoons> "CONST Min ((\<lambda>x. f) ` A)" |
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents:
68795
diff
changeset
|
477 |
"MAX x y. f" \<rightleftharpoons> "MAX x. MAX y. f" |
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents:
68795
diff
changeset
|
478 |
"MAX x. f" \<rightleftharpoons> "CONST Max (CONST range (\<lambda>x. f))" |
9ca183045102
simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents:
68795
diff
changeset
|
479 |
"MAX x\<in>A. f" \<rightleftharpoons> "CONST Max ((\<lambda>x. f) ` A)" |
67969
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
480 |
|
69593 | 481 |
text \<open>An aside: \<^const>\<open>Min\<close>/\<^const>\<open>Max\<close> on linear orders as special case of \<^const>\<open>Inf_fin\<close>/\<^const>\<open>Sup_fin\<close>\<close> |
54864
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
482 |
|
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
483 |
lemma Inf_fin_Min: |
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
484 |
"Inf_fin = (Min :: 'a::{semilattice_inf, linorder} set \<Rightarrow> 'a)" |
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
485 |
by (simp add: Inf_fin_def Min_def inf_min) |
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
486 |
|
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
487 |
lemma Sup_fin_Max: |
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
488 |
"Sup_fin = (Max :: 'a::{semilattice_sup, linorder} set \<Rightarrow> 'a)" |
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
489 |
by (simp add: Sup_fin_def Max_def sup_max) |
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
490 |
|
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
491 |
context linorder |
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
492 |
begin |
a064732223ad
abolished slightly odd global lattice interpretation for min/max
haftmann
parents:
54863
diff
changeset
|
493 |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
494 |
lemma dual_min: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
495 |
"ord.min greater_eq = max" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
496 |
by (auto simp add: ord.min_def max_def fun_eq_iff) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
497 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
498 |
lemma dual_max: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
499 |
"ord.max greater_eq = min" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
500 |
by (auto simp add: ord.max_def min_def fun_eq_iff) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
501 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
502 |
lemma dual_Min: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
503 |
"linorder.Min greater_eq = Max" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
504 |
proof - |
61605 | 505 |
interpret dual: linorder greater_eq greater by (fact dual_linorder) |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
506 |
show ?thesis by (simp add: dual.Min_def dual_min Max_def) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
507 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
508 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
509 |
lemma dual_Max: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
510 |
"linorder.Max greater_eq = Min" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
511 |
proof - |
61605 | 512 |
interpret dual: linorder greater_eq greater by (fact dual_linorder) |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
513 |
show ?thesis by (simp add: dual.Max_def dual_max Min_def) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
514 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
515 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
516 |
lemmas Min_singleton = Min.singleton |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
517 |
lemmas Max_singleton = Max.singleton |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
518 |
lemmas Min_insert = Min.insert |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
519 |
lemmas Max_insert = Max.insert |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
520 |
lemmas Min_Un = Min.union |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
521 |
lemmas Max_Un = Max.union |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
522 |
lemmas hom_Min_commute = Min.hom_commute |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
523 |
lemmas hom_Max_commute = Max.hom_commute |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
524 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
525 |
lemma Min_in [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
526 |
assumes "finite A" and "A \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
527 |
shows "Min A \<in> A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
528 |
using assms by (auto simp add: min_def Min.closed) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
529 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
530 |
lemma Max_in [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
531 |
assumes "finite A" and "A \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
532 |
shows "Max A \<in> A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
533 |
using assms by (auto simp add: max_def Max.closed) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
534 |
|
58467 | 535 |
lemma Min_insert2: |
536 |
assumes "finite A" and min: "\<And>b. b \<in> A \<Longrightarrow> a \<le> b" |
|
537 |
shows "Min (insert a A) = a" |
|
538 |
proof (cases "A = {}") |
|
63915 | 539 |
case True |
540 |
then show ?thesis by simp |
|
58467 | 541 |
next |
63915 | 542 |
case False |
543 |
with \<open>finite A\<close> have "Min (insert a A) = min a (Min A)" |
|
58467 | 544 |
by simp |
60758 | 545 |
moreover from \<open>finite A\<close> \<open>A \<noteq> {}\<close> min have "a \<le> Min A" by simp |
58467 | 546 |
ultimately show ?thesis by (simp add: min.absorb1) |
547 |
qed |
|
548 |
||
549 |
lemma Max_insert2: |
|
550 |
assumes "finite A" and max: "\<And>b. b \<in> A \<Longrightarrow> b \<le> a" |
|
551 |
shows "Max (insert a A) = a" |
|
552 |
proof (cases "A = {}") |
|
63915 | 553 |
case True |
554 |
then show ?thesis by simp |
|
58467 | 555 |
next |
63915 | 556 |
case False |
557 |
with \<open>finite A\<close> have "Max (insert a A) = max a (Max A)" |
|
58467 | 558 |
by simp |
60758 | 559 |
moreover from \<open>finite A\<close> \<open>A \<noteq> {}\<close> max have "Max A \<le> a" by simp |
58467 | 560 |
ultimately show ?thesis by (simp add: max.absorb1) |
561 |
qed |
|
562 |
||
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
563 |
lemma Min_le [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
564 |
assumes "finite A" and "x \<in> A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
565 |
shows "Min A \<le> x" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
566 |
using assms by (fact Min.coboundedI) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
567 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
568 |
lemma Max_ge [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
569 |
assumes "finite A" and "x \<in> A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
570 |
shows "x \<le> Max A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
571 |
using assms by (fact Max.coboundedI) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
572 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
573 |
lemma Min_eqI: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
574 |
assumes "finite A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
575 |
assumes "\<And>y. y \<in> A \<Longrightarrow> y \<ge> x" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
576 |
and "x \<in> A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
577 |
shows "Min A = x" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
578 |
proof (rule antisym) |
60758 | 579 |
from \<open>x \<in> A\<close> have "A \<noteq> {}" by auto |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
580 |
with assms show "Min A \<ge> x" by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
581 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
582 |
from assms show "x \<ge> Min A" by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
583 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
584 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
585 |
lemma Max_eqI: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
586 |
assumes "finite A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
587 |
assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
588 |
and "x \<in> A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
589 |
shows "Max A = x" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
590 |
proof (rule antisym) |
60758 | 591 |
from \<open>x \<in> A\<close> have "A \<noteq> {}" by auto |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
592 |
with assms show "Max A \<le> x" by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
593 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
594 |
from assms show "x \<le> Max A" by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
595 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
596 |
|
66425 | 597 |
lemma eq_Min_iff: |
598 |
"\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> m = Min A \<longleftrightarrow> m \<in> A \<and> (\<forall>a \<in> A. m \<le> a)" |
|
599 |
by (meson Min_in Min_le antisym) |
|
600 |
||
601 |
lemma Min_eq_iff: |
|
602 |
"\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> Min A = m \<longleftrightarrow> m \<in> A \<and> (\<forall>a \<in> A. m \<le> a)" |
|
603 |
by (meson Min_in Min_le antisym) |
|
604 |
||
605 |
lemma eq_Max_iff: |
|
606 |
"\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> m = Max A \<longleftrightarrow> m \<in> A \<and> (\<forall>a \<in> A. a \<le> m)" |
|
607 |
by (meson Max_in Max_ge antisym) |
|
608 |
||
609 |
lemma Max_eq_iff: |
|
610 |
"\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> Max A = m \<longleftrightarrow> m \<in> A \<and> (\<forall>a \<in> A. a \<le> m)" |
|
611 |
by (meson Max_in Max_ge antisym) |
|
612 |
||
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
613 |
context |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
614 |
fixes A :: "'a set" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
615 |
assumes fin_nonempty: "finite A" "A \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
616 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
617 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
618 |
lemma Min_ge_iff [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
619 |
"x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
620 |
using fin_nonempty by (fact Min.bounded_iff) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
621 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
622 |
lemma Max_le_iff [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
623 |
"Max A \<le> x \<longleftrightarrow> (\<forall>a\<in>A. a \<le> x)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
624 |
using fin_nonempty by (fact Max.bounded_iff) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
625 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
626 |
lemma Min_gr_iff [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
627 |
"x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
628 |
using fin_nonempty by (induct rule: finite_ne_induct) simp_all |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
629 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
630 |
lemma Max_less_iff [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
631 |
"Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
632 |
using fin_nonempty by (induct rule: finite_ne_induct) simp_all |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
633 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
634 |
lemma Min_le_iff: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
635 |
"Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
636 |
using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_le_iff_disj) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
637 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
638 |
lemma Max_ge_iff: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
639 |
"x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
640 |
using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: le_max_iff_disj) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
641 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
642 |
lemma Min_less_iff: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
643 |
"Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
644 |
using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_less_iff_disj) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
645 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
646 |
lemma Max_gr_iff: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
647 |
"x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
648 |
using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: less_max_iff_disj) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
649 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
650 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
651 |
|
57800 | 652 |
lemma Max_eq_if: |
653 |
assumes "finite A" "finite B" "\<forall>a\<in>A. \<exists>b\<in>B. a \<le> b" "\<forall>b\<in>B. \<exists>a\<in>A. b \<le> a" |
|
654 |
shows "Max A = Max B" |
|
655 |
proof cases |
|
656 |
assume "A = {}" thus ?thesis using assms by simp |
|
657 |
next |
|
658 |
assume "A \<noteq> {}" thus ?thesis using assms |
|
659 |
by(blast intro: antisym Max_in Max_ge_iff[THEN iffD2]) |
|
660 |
qed |
|
661 |
||
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
662 |
lemma Min_antimono: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
663 |
assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
664 |
shows "Min N \<le> Min M" |
67525
5d04d7bcd5f6
avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents:
67169
diff
changeset
|
665 |
using assms by (fact Min.subset_imp) |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
666 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
667 |
lemma Max_mono: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
668 |
assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
669 |
shows "Max M \<le> Max N" |
67525
5d04d7bcd5f6
avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents:
67169
diff
changeset
|
670 |
using assms by (fact Max.subset_imp) |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
671 |
|
56140
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
wenzelm
parents:
55803
diff
changeset
|
672 |
end |
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
wenzelm
parents:
55803
diff
changeset
|
673 |
|
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
wenzelm
parents:
55803
diff
changeset
|
674 |
context linorder (* FIXME *) |
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
wenzelm
parents:
55803
diff
changeset
|
675 |
begin |
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
wenzelm
parents:
55803
diff
changeset
|
676 |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
677 |
lemma mono_Min_commute: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
678 |
assumes "mono f" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
679 |
assumes "finite A" and "A \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
680 |
shows "f (Min A) = Min (f ` A)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
681 |
proof (rule linorder_class.Min_eqI [symmetric]) |
60758 | 682 |
from \<open>finite A\<close> show "finite (f ` A)" by simp |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
683 |
from assms show "f (Min A) \<in> f ` A" by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
684 |
fix x |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
685 |
assume "x \<in> f ` A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
686 |
then obtain y where "y \<in> A" and "x = f y" .. |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
687 |
with assms have "Min A \<le> y" by auto |
60758 | 688 |
with \<open>mono f\<close> have "f (Min A) \<le> f y" by (rule monoE) |
689 |
with \<open>x = f y\<close> show "f (Min A) \<le> x" by simp |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
690 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
691 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
692 |
lemma mono_Max_commute: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
693 |
assumes "mono f" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
694 |
assumes "finite A" and "A \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
695 |
shows "f (Max A) = Max (f ` A)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
696 |
proof (rule linorder_class.Max_eqI [symmetric]) |
60758 | 697 |
from \<open>finite A\<close> show "finite (f ` A)" by simp |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
698 |
from assms show "f (Max A) \<in> f ` A" by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
699 |
fix x |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
700 |
assume "x \<in> f ` A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
701 |
then obtain y where "y \<in> A" and "x = f y" .. |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
702 |
with assms have "y \<le> Max A" by auto |
60758 | 703 |
with \<open>mono f\<close> have "f y \<le> f (Max A)" by (rule monoE) |
704 |
with \<open>x = f y\<close> show "x \<le> f (Max A)" by simp |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
705 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
706 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
707 |
lemma finite_linorder_max_induct [consumes 1, case_names empty insert]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
708 |
assumes fin: "finite A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
709 |
and empty: "P {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
710 |
and insert: "\<And>b A. finite A \<Longrightarrow> \<forall>a\<in>A. a < b \<Longrightarrow> P A \<Longrightarrow> P (insert b A)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
711 |
shows "P A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
712 |
using fin empty insert |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
713 |
proof (induct rule: finite_psubset_induct) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
714 |
case (psubset A) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
715 |
have IH: "\<And>B. \<lbrakk>B < A; P {}; (\<And>A b. \<lbrakk>finite A; \<forall>a\<in>A. a<b; P A\<rbrakk> \<Longrightarrow> P (insert b A))\<rbrakk> \<Longrightarrow> P B" by fact |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
716 |
have fin: "finite A" by fact |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
717 |
have empty: "P {}" by fact |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
718 |
have step: "\<And>b A. \<lbrakk>finite A; \<forall>a\<in>A. a < b; P A\<rbrakk> \<Longrightarrow> P (insert b A)" by fact |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
719 |
show "P A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
720 |
proof (cases "A = {}") |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
721 |
assume "A = {}" |
60758 | 722 |
then show "P A" using \<open>P {}\<close> by simp |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
723 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
724 |
let ?B = "A - {Max A}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
725 |
let ?A = "insert (Max A) ?B" |
60758 | 726 |
have "finite ?B" using \<open>finite A\<close> by simp |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
727 |
assume "A \<noteq> {}" |
67613 | 728 |
with \<open>finite A\<close> have "Max A \<in> A" by auto |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
729 |
then have A: "?A = A" using insert_Diff_single insert_absorb by auto |
60758 | 730 |
then have "P ?B" using \<open>P {}\<close> step IH [of ?B] by blast |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
731 |
moreover |
60758 | 732 |
have "\<forall>a\<in>?B. a < Max A" using Max_ge [OF \<open>finite A\<close>] by fastforce |
733 |
ultimately show "P A" using A insert_Diff_single step [OF \<open>finite ?B\<close>] by fastforce |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
734 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
735 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
736 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
737 |
lemma finite_linorder_min_induct [consumes 1, case_names empty insert]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
738 |
"\<lbrakk>finite A; P {}; \<And>b A. \<lbrakk>finite A; \<forall>a\<in>A. b < a; P A\<rbrakk> \<Longrightarrow> P (insert b A)\<rbrakk> \<Longrightarrow> P A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
739 |
by (rule linorder.finite_linorder_max_induct [OF dual_linorder]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
740 |
|
70184 | 741 |
lemma finite_ranking_induct[consumes 1, case_names empty insert]: |
742 |
fixes f :: "'b \<Rightarrow> 'a" |
|
743 |
assumes "finite S" |
|
744 |
assumes "P {}" |
|
745 |
assumes "\<And>x S. finite S \<Longrightarrow> (\<And>y. y \<in> S \<Longrightarrow> f y \<le> f x) \<Longrightarrow> P S \<Longrightarrow> P (insert x S)" |
|
746 |
shows "P S" |
|
747 |
using `finite S` |
|
748 |
proof (induction rule: finite_psubset_induct) |
|
749 |
case (psubset A) |
|
750 |
{ |
|
751 |
assume "A \<noteq> {}" |
|
752 |
hence "f ` A \<noteq> {}" and "finite (f ` A)" |
|
753 |
using psubset finite_image_iff by simp+ |
|
754 |
then obtain a where "f a = Max (f ` A)" and "a \<in> A" |
|
755 |
by (metis Max_in[of "f ` A"] imageE) |
|
756 |
then have "P (A - {a})" |
|
757 |
using psubset member_remove by blast |
|
758 |
moreover |
|
759 |
have "\<And>y. y \<in> A \<Longrightarrow> f y \<le> f a" |
|
760 |
using \<open>f a = Max (f ` A)\<close> \<open>finite (f ` A)\<close> by simp |
|
761 |
ultimately |
|
762 |
have ?case |
|
763 |
by (metis \<open>a \<in> A\<close> DiffD1 insert_Diff assms(3) finite_Diff psubset.hyps) |
|
764 |
} |
|
765 |
thus ?case |
|
766 |
using assms(2) by blast |
|
767 |
qed |
|
768 |
||
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
769 |
lemma Least_Min: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
770 |
assumes "finite {a. P a}" and "\<exists>a. P a" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
771 |
shows "(LEAST a. P a) = Min {a. P a}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
772 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
773 |
{ fix A :: "'a set" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
774 |
assume A: "finite A" "A \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
775 |
have "(LEAST a. a \<in> A) = Min A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
776 |
using A proof (induct A rule: finite_ne_induct) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
777 |
case singleton show ?case by (rule Least_equality) simp_all |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
778 |
next |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
779 |
case (insert a A) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
780 |
have "(LEAST b. b = a \<or> b \<in> A) = min a (LEAST a. a \<in> A)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
781 |
by (auto intro!: Least_equality simp add: min_def not_le Min_le_iff insert.hyps dest!: less_imp_le) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
782 |
with insert show ?case by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
783 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
784 |
} from this [of "{a. P a}"] assms show ?thesis by simp |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
785 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
786 |
|
59000 | 787 |
lemma infinite_growing: |
788 |
assumes "X \<noteq> {}" |
|
789 |
assumes *: "\<And>x. x \<in> X \<Longrightarrow> \<exists>y\<in>X. y > x" |
|
790 |
shows "\<not> finite X" |
|
791 |
proof |
|
792 |
assume "finite X" |
|
60758 | 793 |
with \<open>X \<noteq> {}\<close> have "Max X \<in> X" "\<forall>x\<in>X. x \<le> Max X" |
59000 | 794 |
by auto |
795 |
with *[of "Max X"] show False |
|
796 |
by auto |
|
797 |
qed |
|
798 |
||
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
799 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
800 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
801 |
context linordered_ab_semigroup_add |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
802 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
803 |
|
68788 | 804 |
lemma Min_add_commute: |
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
805 |
fixes k |
68793 | 806 |
assumes "finite S" and "S \<noteq> {}" |
807 |
shows "Min ((\<lambda>x. f x + k) ` S) = Min(f ` S) + k" |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
808 |
proof - |
68793 | 809 |
have m: "\<And>x y. min x y + k = min (x+k) (y+k)" |
810 |
by(simp add: min_def antisym add_right_mono) |
|
811 |
have "(\<lambda>x. f x + k) ` S = (\<lambda>y. y + k) ` (f ` S)" by auto |
|
812 |
also have "Min \<dots> = Min (f ` S) + k" |
|
813 |
using assms hom_Min_commute [of "\<lambda>y. y+k" "f ` S", OF m, symmetric] by simp |
|
814 |
finally show ?thesis by simp |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
815 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
816 |
|
68779 | 817 |
lemma Max_add_commute: |
818 |
fixes k |
|
68793 | 819 |
assumes "finite S" and "S \<noteq> {}" |
820 |
shows "Max ((\<lambda>x. f x + k) ` S) = Max(f ` S) + k" |
|
68779 | 821 |
proof - |
68793 | 822 |
have m: "\<And>x y. max x y + k = max (x+k) (y+k)" |
823 |
by(simp add: max_def antisym add_right_mono) |
|
824 |
have "(\<lambda>x. f x + k) ` S = (\<lambda>y. y + k) ` (f ` S)" by auto |
|
825 |
also have "Max \<dots> = Max (f ` S) + k" |
|
826 |
using assms hom_Max_commute [of "\<lambda>y. y+k" "f ` S", OF m, symmetric] by simp |
|
827 |
finally show ?thesis by simp |
|
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
828 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
829 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
830 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
831 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
832 |
context linordered_ab_group_add |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
833 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
834 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
835 |
lemma minus_Max_eq_Min [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
836 |
"finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - Max S = Min (uminus ` S)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
837 |
by (induct S rule: finite_ne_induct) (simp_all add: minus_max_eq_min) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
838 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
839 |
lemma minus_Min_eq_Max [simp]: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
840 |
"finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - Min S = Max (uminus ` S)" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
841 |
by (induct S rule: finite_ne_induct) (simp_all add: minus_min_eq_max) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
842 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
843 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
844 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
845 |
context complete_linorder |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
846 |
begin |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
847 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
848 |
lemma Min_Inf: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
849 |
assumes "finite A" and "A \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
850 |
shows "Min A = Inf A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
851 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
852 |
from assms obtain b B where "A = insert b B" and "finite B" by auto |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
853 |
then show ?thesis |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
854 |
by (simp add: Min.eq_fold complete_linorder_inf_min [symmetric] inf_Inf_fold_inf inf.commute [of b]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
855 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
856 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
857 |
lemma Max_Sup: |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
858 |
assumes "finite A" and "A \<noteq> {}" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
859 |
shows "Max A = Sup A" |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
860 |
proof - |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
861 |
from assms obtain b B where "A = insert b B" and "finite B" by auto |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
862 |
then show ?thesis |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
863 |
by (simp add: Max.eq_fold complete_linorder_sup_max [symmetric] sup_Sup_fold_sup sup.commute [of b]) |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
864 |
qed |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
865 |
|
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
866 |
end |
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
867 |
|
65817 | 868 |
|
869 |
subsection \<open>Arg Min\<close> |
|
870 |
||
65842 | 871 |
definition is_arg_min :: "('a \<Rightarrow> 'b::ord) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" where |
872 |
"is_arg_min f P x = (P x \<and> \<not>(\<exists>y. P y \<and> f y < f x))" |
|
873 |
||
874 |
definition arg_min :: "('a \<Rightarrow> 'b::ord) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a" where |
|
875 |
"arg_min f P = (SOME x. is_arg_min f P x)" |
|
876 |
||
67169 | 877 |
definition arg_min_on :: "('a \<Rightarrow> 'b::ord) \<Rightarrow> 'a set \<Rightarrow> 'a" where |
878 |
"arg_min_on f S = arg_min f (\<lambda>x. x \<in> S)" |
|
65842 | 879 |
|
65951 | 880 |
syntax |
881 |
"_arg_min" :: "('a \<Rightarrow> 'b) \<Rightarrow> pttrn \<Rightarrow> bool \<Rightarrow> 'a" |
|
67036 | 882 |
("(3ARG'_MIN _ _./ _)" [1000, 0, 10] 10) |
65951 | 883 |
translations |
884 |
"ARG_MIN f x. P" \<rightleftharpoons> "CONST arg_min f (\<lambda>x. P)" |
|
885 |
||
65842 | 886 |
lemma is_arg_min_linorder: fixes f :: "'a \<Rightarrow> 'b :: linorder" |
887 |
shows "is_arg_min f P x = (P x \<and> (\<forall>y. P y \<longrightarrow> f x \<le> f y))" |
|
888 |
by(auto simp add: is_arg_min_def) |
|
65817 | 889 |
|
67566 | 890 |
lemma is_arg_min_antimono: fixes f :: "'a \<Rightarrow> ('b::order)" |
891 |
shows "\<lbrakk> is_arg_min f P x; f y \<le> f x; P y \<rbrakk> \<Longrightarrow> is_arg_min f P y" |
|
892 |
by (simp add: order.order_iff_strict is_arg_min_def) |
|
893 |
||
65951 | 894 |
lemma arg_minI: |
895 |
"\<lbrakk> P x; |
|
896 |
\<And>y. P y \<Longrightarrow> \<not> f y < f x; |
|
897 |
\<And>x. \<lbrakk> P x; \<forall>y. P y \<longrightarrow> \<not> f y < f x \<rbrakk> \<Longrightarrow> Q x \<rbrakk> |
|
898 |
\<Longrightarrow> Q (arg_min f P)" |
|
899 |
apply (simp add: arg_min_def is_arg_min_def) |
|
900 |
apply (rule someI2_ex) |
|
901 |
apply blast |
|
902 |
apply blast |
|
903 |
done |
|
904 |
||
905 |
lemma arg_min_equality: |
|
65952 | 906 |
"\<lbrakk> P k; \<And>x. P x \<Longrightarrow> f k \<le> f x \<rbrakk> \<Longrightarrow> f (arg_min f P) = f k" |
65951 | 907 |
for f :: "_ \<Rightarrow> 'a::order" |
908 |
apply (rule arg_minI) |
|
909 |
apply assumption |
|
910 |
apply (simp add: less_le_not_le) |
|
911 |
by (metis le_less) |
|
912 |
||
65952 | 913 |
lemma wf_linord_ex_has_least: |
914 |
"\<lbrakk> wf r; \<forall>x y. (x, y) \<in> r\<^sup>+ \<longleftrightarrow> (y, x) \<notin> r\<^sup>*; P k \<rbrakk> |
|
915 |
\<Longrightarrow> \<exists>x. P x \<and> (\<forall>y. P y \<longrightarrow> (m x, m y) \<in> r\<^sup>*)" |
|
916 |
apply (drule wf_trancl [THEN wf_eq_minimal [THEN iffD1]]) |
|
917 |
apply (drule_tac x = "m ` Collect P" in spec) |
|
918 |
by force |
|
919 |
||
920 |
lemma ex_has_least_nat: "P k \<Longrightarrow> \<exists>x. P x \<and> (\<forall>y. P y \<longrightarrow> m x \<le> m y)" |
|
921 |
for m :: "'a \<Rightarrow> nat" |
|
922 |
apply (simp only: pred_nat_trancl_eq_le [symmetric]) |
|
923 |
apply (rule wf_pred_nat [THEN wf_linord_ex_has_least]) |
|
924 |
apply (simp add: less_eq linorder_not_le pred_nat_trancl_eq_le) |
|
925 |
by assumption |
|
926 |
||
65951 | 927 |
lemma arg_min_nat_lemma: |
928 |
"P k \<Longrightarrow> P(arg_min m P) \<and> (\<forall>y. P y \<longrightarrow> m (arg_min m P) \<le> m y)" |
|
65842 | 929 |
for m :: "'a \<Rightarrow> nat" |
930 |
apply (simp add: arg_min_def is_arg_min_linorder) |
|
931 |
apply (rule someI_ex) |
|
932 |
apply (erule ex_has_least_nat) |
|
933 |
done |
|
934 |
||
935 |
lemmas arg_min_natI = arg_min_nat_lemma [THEN conjunct1] |
|
65817 | 936 |
|
65951 | 937 |
lemma is_arg_min_arg_min_nat: fixes m :: "'a \<Rightarrow> nat" |
938 |
shows "P x \<Longrightarrow> is_arg_min m P (arg_min m P)" |
|
939 |
by (metis arg_min_nat_lemma is_arg_min_linorder) |
|
940 |
||
65842 | 941 |
lemma arg_min_nat_le: "P x \<Longrightarrow> m (arg_min m P) \<le> m x" |
942 |
for m :: "'a \<Rightarrow> nat" |
|
943 |
by (rule arg_min_nat_lemma [THEN conjunct2, THEN spec, THEN mp]) |
|
944 |
||
945 |
lemma ex_min_if_finite: |
|
946 |
"\<lbrakk> finite S; S \<noteq> {} \<rbrakk> \<Longrightarrow> \<exists>m\<in>S. \<not>(\<exists>x\<in>S. x < (m::'a::order))" |
|
947 |
by(induction rule: finite.induct) (auto intro: order.strict_trans) |
|
948 |
||
949 |
lemma ex_is_arg_min_if_finite: fixes f :: "'a \<Rightarrow> 'b :: order" |
|
67613 | 950 |
shows "\<lbrakk> finite S; S \<noteq> {} \<rbrakk> \<Longrightarrow> \<exists>x. is_arg_min f (\<lambda>x. x \<in> S) x" |
65842 | 951 |
unfolding is_arg_min_def |
952 |
using ex_min_if_finite[of "f ` S"] |
|
953 |
by auto |
|
65817 | 954 |
|
955 |
lemma arg_min_SOME_Min: |
|
65842 | 956 |
"finite S \<Longrightarrow> arg_min_on f S = (SOME y. y \<in> S \<and> f y = Min(f ` S))" |
67169 | 957 |
unfolding arg_min_on_def arg_min_def is_arg_min_linorder |
65817 | 958 |
apply(rule arg_cong[where f = Eps]) |
959 |
apply (auto simp: fun_eq_iff intro: Min_eqI[symmetric]) |
|
960 |
done |
|
961 |
||
65842 | 962 |
lemma arg_min_if_finite: fixes f :: "'a \<Rightarrow> 'b :: order" |
963 |
assumes "finite S" "S \<noteq> {}" |
|
964 |
shows "arg_min_on f S \<in> S" and "\<not>(\<exists>x\<in>S. f x < f (arg_min_on f S))" |
|
965 |
using ex_is_arg_min_if_finite[OF assms, of f] |
|
67169 | 966 |
unfolding arg_min_on_def arg_min_def is_arg_min_def |
65842 | 967 |
by(auto dest!: someI_ex) |
65817 | 968 |
|
969 |
lemma arg_min_least: fixes f :: "'a \<Rightarrow> 'b :: linorder" |
|
65842 | 970 |
shows "\<lbrakk> finite S; S \<noteq> {}; y \<in> S \<rbrakk> \<Longrightarrow> f(arg_min_on f S) \<le> f y" |
65817 | 971 |
by(simp add: arg_min_SOME_Min inv_into_def2[symmetric] f_inv_into_f) |
972 |
||
973 |
lemma arg_min_inj_eq: fixes f :: "'a \<Rightarrow> 'b :: order" |
|
65842 | 974 |
shows "\<lbrakk> inj_on f {x. P x}; P a; \<forall>y. P y \<longrightarrow> f a \<le> f y \<rbrakk> \<Longrightarrow> arg_min f P = a" |
975 |
apply(simp add: arg_min_def is_arg_min_def) |
|
65817 | 976 |
apply(rule someI2[of _ a]) |
977 |
apply (simp add: less_le_not_le) |
|
65842 | 978 |
by (metis inj_on_eq_iff less_le mem_Collect_eq) |
65817 | 979 |
|
65954 | 980 |
|
981 |
subsection \<open>Arg Max\<close> |
|
982 |
||
983 |
definition is_arg_max :: "('a \<Rightarrow> 'b::ord) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" where |
|
984 |
"is_arg_max f P x = (P x \<and> \<not>(\<exists>y. P y \<and> f y > f x))" |
|
985 |
||
986 |
definition arg_max :: "('a \<Rightarrow> 'b::ord) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a" where |
|
987 |
"arg_max f P = (SOME x. is_arg_max f P x)" |
|
988 |
||
67169 | 989 |
definition arg_max_on :: "('a \<Rightarrow> 'b::ord) \<Rightarrow> 'a set \<Rightarrow> 'a" where |
990 |
"arg_max_on f S = arg_max f (\<lambda>x. x \<in> S)" |
|
65954 | 991 |
|
992 |
syntax |
|
993 |
"_arg_max" :: "('a \<Rightarrow> 'b) \<Rightarrow> pttrn \<Rightarrow> bool \<Rightarrow> 'a" |
|
67036 | 994 |
("(3ARG'_MAX _ _./ _)" [1000, 0, 10] 10) |
65954 | 995 |
translations |
996 |
"ARG_MAX f x. P" \<rightleftharpoons> "CONST arg_max f (\<lambda>x. P)" |
|
997 |
||
998 |
lemma is_arg_max_linorder: fixes f :: "'a \<Rightarrow> 'b :: linorder" |
|
999 |
shows "is_arg_max f P x = (P x \<and> (\<forall>y. P y \<longrightarrow> f x \<ge> f y))" |
|
1000 |
by(auto simp add: is_arg_max_def) |
|
1001 |
||
1002 |
lemma arg_maxI: |
|
1003 |
"P x \<Longrightarrow> |
|
1004 |
(\<And>y. P y \<Longrightarrow> \<not> f y > f x) \<Longrightarrow> |
|
1005 |
(\<And>x. P x \<Longrightarrow> \<forall>y. P y \<longrightarrow> \<not> f y > f x \<Longrightarrow> Q x) \<Longrightarrow> |
|
1006 |
Q (arg_max f P)" |
|
1007 |
apply (simp add: arg_max_def is_arg_max_def) |
|
1008 |
apply (rule someI2_ex) |
|
1009 |
apply blast |
|
1010 |
apply blast |
|
1011 |
done |
|
1012 |
||
1013 |
lemma arg_max_equality: |
|
1014 |
"\<lbrakk> P k; \<And>x. P x \<Longrightarrow> f x \<le> f k \<rbrakk> \<Longrightarrow> f (arg_max f P) = f k" |
|
1015 |
for f :: "_ \<Rightarrow> 'a::order" |
|
1016 |
apply (rule arg_maxI [where f = f]) |
|
1017 |
apply assumption |
|
1018 |
apply (simp add: less_le_not_le) |
|
1019 |
by (metis le_less) |
|
1020 |
||
1021 |
lemma ex_has_greatest_nat_lemma: |
|
1022 |
"P k \<Longrightarrow> \<forall>x. P x \<longrightarrow> (\<exists>y. P y \<and> \<not> f y \<le> f x) \<Longrightarrow> \<exists>y. P y \<and> \<not> f y < f k + n" |
|
1023 |
for f :: "'a \<Rightarrow> nat" |
|
1024 |
by (induct n) (force simp: le_Suc_eq)+ |
|
1025 |
||
1026 |
lemma ex_has_greatest_nat: |
|
1027 |
"P k \<Longrightarrow> \<forall>y. P y \<longrightarrow> f y < b \<Longrightarrow> \<exists>x. P x \<and> (\<forall>y. P y \<longrightarrow> f y \<le> f x)" |
|
1028 |
for f :: "'a \<Rightarrow> nat" |
|
1029 |
apply (rule ccontr) |
|
1030 |
apply (cut_tac P = P and n = "b - f k" in ex_has_greatest_nat_lemma) |
|
1031 |
apply (subgoal_tac [3] "f k \<le> b") |
|
1032 |
apply auto |
|
1033 |
done |
|
1034 |
||
1035 |
lemma arg_max_nat_lemma: |
|
1036 |
"\<lbrakk> P k; \<forall>y. P y \<longrightarrow> f y < b \<rbrakk> |
|
1037 |
\<Longrightarrow> P (arg_max f P) \<and> (\<forall>y. P y \<longrightarrow> f y \<le> f (arg_max f P))" |
|
1038 |
for f :: "'a \<Rightarrow> nat" |
|
1039 |
apply (simp add: arg_max_def is_arg_max_linorder) |
|
1040 |
apply (rule someI_ex) |
|
1041 |
apply (erule (1) ex_has_greatest_nat) |
|
1042 |
done |
|
1043 |
||
1044 |
lemmas arg_max_natI = arg_max_nat_lemma [THEN conjunct1] |
|
1045 |
||
1046 |
lemma arg_max_nat_le: "P x \<Longrightarrow> \<forall>y. P y \<longrightarrow> f y < b \<Longrightarrow> f x \<le> f (arg_max f P)" |
|
1047 |
for f :: "'a \<Rightarrow> nat" |
|
1048 |
by (blast dest: arg_max_nat_lemma [THEN conjunct2, THEN spec, of P]) |
|
1049 |
||
54744
1e7f2d296e19
more algebraic terminology for theories about big operators
haftmann
parents:
diff
changeset
|
1050 |
end |