| author | hoelzl | 
| Wed, 10 Oct 2012 12:12:27 +0200 | |
| changeset 49789 | e0a4cb91a8a9 | 
| parent 49715 | 16d8c6d288bc | 
| child 49962 | a8cc904a6820 | 
| permissions | -rw-r--r-- | 
| 
35719
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1  | 
(* Title: HOL/Big_Operators.thy  | 
| 12396 | 2  | 
Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel  | 
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
3  | 
with contributions by Jeremy Avigad  | 
| 12396 | 4  | 
*)  | 
5  | 
||
| 
35719
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
6  | 
header {* Big operators and finite (non-empty) sets *}
 | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
7  | 
|
| 
35719
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
8  | 
theory Big_Operators  | 
| 
35722
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
9  | 
imports Plain  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
10  | 
begin  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
11  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
12  | 
subsection {* Generic monoid operation over a set *}
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
13  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
14  | 
no_notation times (infixl "*" 70)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
15  | 
no_notation Groups.one ("1")
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
16  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
17  | 
locale comm_monoid_big = comm_monoid +  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
18  | 
  fixes F :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
19  | 
assumes F_eq: "F g A = (if finite A then fold_image (op *) g 1 A else 1)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
20  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
21  | 
sublocale comm_monoid_big < folding_image proof  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
22  | 
qed (simp add: F_eq)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
23  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
24  | 
context comm_monoid_big  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
25  | 
begin  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
26  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
27  | 
lemma infinite [simp]:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
28  | 
"\<not> finite A \<Longrightarrow> F g A = 1"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
29  | 
by (simp add: F_eq)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
30  | 
|
| 42986 | 31  | 
lemma F_cong:  | 
32  | 
assumes "A = B" "\<And>x. x \<in> B \<Longrightarrow> h x = g x"  | 
|
33  | 
shows "F h A = F g B"  | 
|
34  | 
proof cases  | 
|
35  | 
assume "finite A"  | 
|
36  | 
with assms show ?thesis unfolding `A = B` by (simp cong: cong)  | 
|
37  | 
next  | 
|
38  | 
assume "\<not> finite A"  | 
|
39  | 
then show ?thesis unfolding `A = B` by simp  | 
|
40  | 
qed  | 
|
41  | 
||
| 48849 | 42  | 
lemma strong_F_cong [cong]:  | 
43  | 
"\<lbrakk> A = B; !!x. x:B =simp=> g x = h x \<rbrakk>  | 
|
44  | 
\<Longrightarrow> F (%x. g x) A = F (%x. h x) B"  | 
|
45  | 
by (rule F_cong) (simp_all add: simp_implies_def)  | 
|
46  | 
||
| 48821 | 47  | 
lemma F_neutral[simp]: "F (%i. 1) A = 1"  | 
48  | 
by (cases "finite A") (simp_all add: neutral)  | 
|
49  | 
||
50  | 
lemma F_neutral': "ALL a:A. g a = 1 \<Longrightarrow> F g A = 1"  | 
|
| 48849 | 51  | 
by simp  | 
52  | 
||
53  | 
lemma F_subset_diff: "\<lbrakk> B \<subseteq> A; finite A \<rbrakk> \<Longrightarrow> F g A = F g (A - B) * F g B"  | 
|
54  | 
by (metis Diff_partition union_disjoint Diff_disjoint finite_Un inf_commute sup_commute)  | 
|
55  | 
||
56  | 
lemma F_mono_neutral_cong_left:  | 
|
57  | 
assumes "finite T" and "S \<subseteq> T" and "\<forall>i \<in> T - S. h i = 1"  | 
|
58  | 
and "\<And>x. x \<in> S \<Longrightarrow> g x = h x" shows "F g S = F h T"  | 
|
59  | 
proof-  | 
|
60  | 
have eq: "T = S \<union> (T - S)" using `S \<subseteq> T` by blast  | 
|
61  | 
  have d: "S \<inter> (T - S) = {}" using `S \<subseteq> T` by blast
 | 
|
62  | 
from `finite T` `S \<subseteq> T` have f: "finite S" "finite (T - S)"  | 
|
63  | 
by (auto intro: finite_subset)  | 
|
64  | 
show ?thesis using assms(4)  | 
|
65  | 
by (simp add: union_disjoint[OF f d, unfolded eq[symmetric]] F_neutral'[OF assms(3)])  | 
|
66  | 
qed  | 
|
67  | 
||
| 48850 | 68  | 
lemma F_mono_neutral_cong_right:  | 
69  | 
"\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 1; \<And>x. x \<in> S \<Longrightarrow> g x = h x \<rbrakk>  | 
|
70  | 
\<Longrightarrow> F g T = F h S"  | 
|
71  | 
by(auto intro!: F_mono_neutral_cong_left[symmetric])  | 
|
| 48849 | 72  | 
|
73  | 
lemma F_mono_neutral_left:  | 
|
74  | 
"\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 1 \<rbrakk> \<Longrightarrow> F g S = F g T"  | 
|
75  | 
by(blast intro: F_mono_neutral_cong_left)  | 
|
76  | 
||
| 48850 | 77  | 
lemma F_mono_neutral_right:  | 
78  | 
"\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 1 \<rbrakk> \<Longrightarrow> F g T = F g S"  | 
|
79  | 
by(blast intro!: F_mono_neutral_left[symmetric])  | 
|
| 48849 | 80  | 
|
81  | 
lemma F_delta:  | 
|
82  | 
assumes fS: "finite S"  | 
|
83  | 
shows "F (\<lambda>k. if k=a then b k else 1) S = (if a \<in> S then b a else 1)"  | 
|
84  | 
proof-  | 
|
85  | 
let ?f = "(\<lambda>k. if k=a then b k else 1)"  | 
|
86  | 
  { assume a: "a \<notin> S"
 | 
|
87  | 
hence "\<forall>k\<in>S. ?f k = 1" by simp  | 
|
88  | 
hence ?thesis using a by simp }  | 
|
89  | 
moreover  | 
|
90  | 
  { assume a: "a \<in> S"
 | 
|
91  | 
    let ?A = "S - {a}"
 | 
|
92  | 
    let ?B = "{a}"
 | 
|
93  | 
have eq: "S = ?A \<union> ?B" using a by blast  | 
|
94  | 
    have dj: "?A \<inter> ?B = {}" by simp
 | 
|
95  | 
from fS have fAB: "finite ?A" "finite ?B" by auto  | 
|
96  | 
have "F ?f S = F ?f ?A * F ?f ?B"  | 
|
97  | 
using union_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]  | 
|
98  | 
by simp  | 
|
99  | 
then have ?thesis using a by simp }  | 
|
100  | 
ultimately show ?thesis by blast  | 
|
101  | 
qed  | 
|
102  | 
||
103  | 
lemma F_delta':  | 
|
104  | 
assumes fS: "finite S" shows  | 
|
105  | 
"F (\<lambda>k. if a = k then b k else 1) S = (if a \<in> S then b a else 1)"  | 
|
106  | 
using F_delta[OF fS, of a b, symmetric] by (auto intro: F_cong)  | 
|
| 48821 | 107  | 
|
| 48861 | 108  | 
lemma F_fun_f: "F (%x. g x * h x) A = (F g A * F h A)"  | 
109  | 
by (cases "finite A") (simp_all add: distrib)  | 
|
110  | 
||
| 48893 | 111  | 
|
112  | 
text {* for ad-hoc proofs for @{const fold_image} *}
 | 
|
113  | 
lemma comm_monoid_mult: "class.comm_monoid_mult (op *) 1"  | 
|
114  | 
proof qed (auto intro: assoc commute)  | 
|
115  | 
||
116  | 
lemma F_Un_neutral:  | 
|
117  | 
assumes fS: "finite S" and fT: "finite T"  | 
|
118  | 
and I1: "\<forall>x \<in> S\<inter>T. g x = 1"  | 
|
119  | 
shows "F g (S \<union> T) = F g S * F g T"  | 
|
120  | 
proof -  | 
|
121  | 
interpret comm_monoid_mult "op *" 1 by (fact comm_monoid_mult)  | 
|
122  | 
show ?thesis  | 
|
123  | 
using fS fT  | 
|
124  | 
apply (simp add: F_eq)  | 
|
125  | 
apply (rule fold_image_Un_one)  | 
|
126  | 
using I1 by auto  | 
|
127  | 
qed  | 
|
128  | 
||
| 42986 | 129  | 
lemma If_cases:  | 
130  | 
fixes P :: "'b \<Rightarrow> bool" and g h :: "'b \<Rightarrow> 'a"  | 
|
131  | 
assumes fA: "finite A"  | 
|
132  | 
shows "F (\<lambda>x. if P x then h x else g x) A =  | 
|
133  | 
         F h (A \<inter> {x. P x}) * F g (A \<inter> - {x. P x})"
 | 
|
134  | 
proof-  | 
|
135  | 
  have a: "A = A \<inter> {x. P x} \<union> A \<inter> -{x. P x}" 
 | 
|
136  | 
          "(A \<inter> {x. P x}) \<inter> (A \<inter> -{x. P x}) = {}" 
 | 
|
137  | 
by blast+  | 
|
138  | 
from fA  | 
|
139  | 
  have f: "finite (A \<inter> {x. P x})" "finite (A \<inter> -{x. P x})" by auto
 | 
|
140  | 
let ?g = "\<lambda>x. if P x then h x else g x"  | 
|
141  | 
from union_disjoint[OF f a(2), of ?g] a(1)  | 
|
142  | 
show ?thesis  | 
|
143  | 
by (subst (1 2) F_cong) simp_all  | 
|
144  | 
qed  | 
|
145  | 
||
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
146  | 
end  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
147  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
148  | 
text {* for ad-hoc proofs for @{const fold_image} *}
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
149  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
150  | 
lemma (in comm_monoid_add) comm_monoid_mult:  | 
| 
36635
 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 
haftmann 
parents: 
36622 
diff
changeset
 | 
151  | 
"class.comm_monoid_mult (op +) 0"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
152  | 
proof qed (auto intro: add_assoc add_commute)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
153  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
154  | 
notation times (infixl "*" 70)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
155  | 
notation Groups.one ("1")
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
156  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
157  | 
|
| 15402 | 158  | 
subsection {* Generalized summation over a set *}
 | 
159  | 
||
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
160  | 
definition (in comm_monoid_add) setsum :: "('b \<Rightarrow> 'a) => 'b set => 'a" where
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
161  | 
"setsum f A = (if finite A then fold_image (op +) f 0 A else 0)"  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
162  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
163  | 
sublocale comm_monoid_add < setsum!: comm_monoid_big "op +" 0 setsum proof  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
164  | 
qed (fact setsum_def)  | 
| 15402 | 165  | 
|
| 19535 | 166  | 
abbreviation  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21249 
diff
changeset
 | 
167  | 
  Setsum  ("\<Sum>_" [1000] 999) where
 | 
| 19535 | 168  | 
"\<Sum>A == setsum (%x. x) A"  | 
169  | 
||
| 15402 | 170  | 
text{* Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is
 | 
171  | 
written @{text"\<Sum>x\<in>A. e"}. *}
 | 
|
172  | 
||
173  | 
syntax  | 
|
| 17189 | 174  | 
  "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add"    ("(3SUM _:_. _)" [0, 51, 10] 10)
 | 
| 15402 | 175  | 
syntax (xsymbols)  | 
| 17189 | 176  | 
  "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add"    ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
 | 
| 15402 | 177  | 
syntax (HTML output)  | 
| 17189 | 178  | 
  "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add"    ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
 | 
| 15402 | 179  | 
|
180  | 
translations -- {* Beware of argument permutation! *}
 | 
|
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
181  | 
"SUM i:A. b" == "CONST setsum (%i. b) A"  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
182  | 
"\<Sum>i\<in>A. b" == "CONST setsum (%i. b) A"  | 
| 15402 | 183  | 
|
184  | 
text{* Instead of @{term"\<Sum>x\<in>{x. P}. e"} we introduce the shorter
 | 
|
185  | 
 @{text"\<Sum>x|P. e"}. *}
 | 
|
186  | 
||
187  | 
syntax  | 
|
| 17189 | 188  | 
  "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10)
 | 
| 15402 | 189  | 
syntax (xsymbols)  | 
| 17189 | 190  | 
  "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
 | 
| 15402 | 191  | 
syntax (HTML output)  | 
| 17189 | 192  | 
  "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
 | 
| 15402 | 193  | 
|
194  | 
translations  | 
|
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
195  | 
  "SUM x|P. t" => "CONST setsum (%x. t) {x. P}"
 | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
196  | 
  "\<Sum>x|P. t" => "CONST setsum (%x. t) {x. P}"
 | 
| 15402 | 197  | 
|
198  | 
print_translation {*
 | 
|
199  | 
let  | 
|
| 35115 | 200  | 
  fun setsum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
 | 
201  | 
if x <> y then raise Match  | 
|
202  | 
else  | 
|
203  | 
let  | 
|
| 
49660
 
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
 
wenzelm 
parents: 
48893 
diff
changeset
 | 
204  | 
val x' = Syntax_Trans.mark_bound_body (x, Tx);  | 
| 35115 | 205  | 
val t' = subst_bound (x', t);  | 
206  | 
val P' = subst_bound (x', P);  | 
|
| 
49660
 
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
 
wenzelm 
parents: 
48893 
diff
changeset
 | 
207  | 
in  | 
| 
 
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
 
wenzelm 
parents: 
48893 
diff
changeset
 | 
208  | 
            Syntax.const @{syntax_const "_qsetsum"} $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
 | 
| 
 
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
 
wenzelm 
parents: 
48893 
diff
changeset
 | 
209  | 
end  | 
| 35115 | 210  | 
| setsum_tr' _ = raise Match;  | 
211  | 
in [(@{const_syntax setsum}, setsum_tr')] end
 | 
|
| 15402 | 212  | 
*}  | 
213  | 
||
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
214  | 
lemma setsum_empty:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
215  | 
  "setsum f {} = 0"
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
216  | 
by (fact setsum.empty)  | 
| 15402 | 217  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
218  | 
lemma setsum_insert:  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
219  | 
"finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
220  | 
by (fact setsum.insert)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
221  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
222  | 
lemma setsum_infinite:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
223  | 
"~ finite A ==> setsum f A = 0"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
224  | 
by (fact setsum.infinite)  | 
| 15402 | 225  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
226  | 
lemma (in comm_monoid_add) setsum_reindex:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
227  | 
assumes "inj_on f B" shows "setsum h (f ` B) = setsum (h \<circ> f) B"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
228  | 
proof -  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
229  | 
interpret comm_monoid_mult "op +" 0 by (fact comm_monoid_mult)  | 
| 48849 | 230  | 
from assms show ?thesis by (auto simp add: setsum_def fold_image_reindex o_def dest!:finite_imageD)  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
231  | 
qed  | 
| 
15409
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
232  | 
|
| 48849 | 233  | 
lemma setsum_reindex_id:  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
234  | 
"inj_on f B ==> setsum f B = setsum id (f ` B)"  | 
| 48849 | 235  | 
by (simp add: setsum_reindex)  | 
| 15402 | 236  | 
|
| 48849 | 237  | 
lemma setsum_reindex_nonzero:  | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
238  | 
assumes fS: "finite S"  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
239  | 
and nz: "\<And> x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x \<noteq> y \<Longrightarrow> f x = f y \<Longrightarrow> h (f x) = 0"  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
240  | 
shows "setsum h (f ` S) = setsum (h o f) S"  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
241  | 
using nz  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
242  | 
proof(induct rule: finite_induct[OF fS])  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
243  | 
case 1 thus ?case by simp  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
244  | 
next  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
245  | 
case (2 x F)  | 
| 48849 | 246  | 
  { assume fxF: "f x \<in> f ` F" hence "\<exists>y \<in> F . f y = f x" by auto
 | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
247  | 
then obtain y where y: "y \<in> F" "f x = f y" by auto  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
248  | 
from "2.hyps" y have xy: "x \<noteq> y" by auto  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
249  | 
|
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
250  | 
from "2.prems"[of x y] "2.hyps" xy y have h0: "h (f x) = 0" by simp  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
251  | 
have "setsum h (f ` insert x F) = setsum h (f ` F)" using fxF by auto  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
252  | 
also have "\<dots> = setsum (h o f) (insert x F)"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
253  | 
unfolding setsum.insert[OF `finite F` `x\<notin>F`]  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
254  | 
using h0  | 
| 48849 | 255  | 
apply (simp cong del:setsum.strong_F_cong)  | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
256  | 
apply (rule "2.hyps"(3))  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
257  | 
apply (rule_tac y="y" in "2.prems")  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
258  | 
apply simp_all  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
259  | 
done  | 
| 48849 | 260  | 
finally have ?case . }  | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
261  | 
moreover  | 
| 48849 | 262  | 
  { assume fxF: "f x \<notin> f ` F"
 | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
263  | 
have "setsum h (f ` insert x F) = h (f x) + setsum h (f ` F)"  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
264  | 
using fxF "2.hyps" by simp  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
265  | 
also have "\<dots> = setsum (h o f) (insert x F)"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
266  | 
unfolding setsum.insert[OF `finite F` `x\<notin>F`]  | 
| 48849 | 267  | 
apply (simp cong del:setsum.strong_F_cong)  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
268  | 
apply (rule cong [OF refl [of "op + (h (f x))"]])  | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
269  | 
apply (rule "2.hyps"(3))  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
270  | 
apply (rule_tac y="y" in "2.prems")  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
271  | 
apply simp_all  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
272  | 
done  | 
| 48849 | 273  | 
finally have ?case . }  | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
274  | 
ultimately show ?case by blast  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
275  | 
qed  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
276  | 
|
| 48849 | 277  | 
lemma setsum_cong:  | 
| 15402 | 278  | 
"A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B"  | 
| 48849 | 279  | 
by (fact setsum.F_cong)  | 
| 15402 | 280  | 
|
| 48849 | 281  | 
lemma strong_setsum_cong:  | 
| 
16733
 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 
nipkow 
parents: 
16632 
diff
changeset
 | 
282  | 
"A = B ==> (!!x. x:B =simp=> f x = g x)  | 
| 
 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 
nipkow 
parents: 
16632 
diff
changeset
 | 
283  | 
==> setsum (%x. f x) A = setsum (%x. g x) B"  | 
| 48849 | 284  | 
by (fact setsum.strong_F_cong)  | 
| 
16632
 
ad2895beef79
Added strong_setsum_cong and strong_setprod_cong.
 
berghofe 
parents: 
16550 
diff
changeset
 | 
285  | 
|
| 48849 | 286  | 
lemma setsum_cong2: "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setsum f A = setsum g A"  | 
287  | 
by (auto intro: setsum_cong)  | 
|
| 15554 | 288  | 
|
| 48849 | 289  | 
lemma setsum_reindex_cong:  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
290  | 
"[|inj_on f A; B = f ` A; !!a. a:A \<Longrightarrow> g a = h (f a)|]  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
291  | 
==> setsum h B = setsum g A"  | 
| 48849 | 292  | 
by (simp add: setsum_reindex)  | 
| 15402 | 293  | 
|
| 48821 | 294  | 
lemmas setsum_0 = setsum.F_neutral  | 
295  | 
lemmas setsum_0' = setsum.F_neutral'  | 
|
| 15402 | 296  | 
|
| 48849 | 297  | 
lemma setsum_Un_Int: "finite A ==> finite B ==>  | 
| 15402 | 298  | 
setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B"  | 
299  | 
  -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
 | 
|
| 48849 | 300  | 
by (fact setsum.union_inter)  | 
| 15402 | 301  | 
|
| 48849 | 302  | 
lemma setsum_Un_disjoint: "finite A ==> finite B  | 
| 15402 | 303  | 
  ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
 | 
| 48849 | 304  | 
by (fact setsum.union_disjoint)  | 
305  | 
||
306  | 
lemma setsum_subset_diff: "\<lbrakk> B \<subseteq> A; finite A \<rbrakk> \<Longrightarrow>  | 
|
307  | 
setsum f A = setsum f (A - B) + setsum f B"  | 
|
308  | 
by(fact setsum.F_subset_diff)  | 
|
| 15402 | 309  | 
|
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
310  | 
lemma setsum_mono_zero_left:  | 
| 48849 | 311  | 
"\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. f i = 0 \<rbrakk> \<Longrightarrow> setsum f S = setsum f T"  | 
312  | 
by(fact setsum.F_mono_neutral_left)  | 
|
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
313  | 
|
| 48849 | 314  | 
lemmas setsum_mono_zero_right = setsum.F_mono_neutral_right  | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
315  | 
|
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
316  | 
lemma setsum_mono_zero_cong_left:  | 
| 48849 | 317  | 
"\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 0; \<And>x. x \<in> S \<Longrightarrow> f x = g x \<rbrakk>  | 
318  | 
\<Longrightarrow> setsum f S = setsum g T"  | 
|
319  | 
by(fact setsum.F_mono_neutral_cong_left)  | 
|
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
320  | 
|
| 48849 | 321  | 
lemmas setsum_mono_zero_cong_right = setsum.F_mono_neutral_cong_right  | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
322  | 
|
| 48849 | 323  | 
lemma setsum_delta: "finite S \<Longrightarrow>  | 
324  | 
setsum (\<lambda>k. if k=a then b k else 0) S = (if a \<in> S then b a else 0)"  | 
|
325  | 
by(fact setsum.F_delta)  | 
|
326  | 
||
327  | 
lemma setsum_delta': "finite S \<Longrightarrow>  | 
|
328  | 
setsum (\<lambda>k. if a = k then b k else 0) S = (if a\<in> S then b a else 0)"  | 
|
329  | 
by(fact setsum.F_delta')  | 
|
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
330  | 
|
| 
30260
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
331  | 
lemma setsum_restrict_set:  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
332  | 
assumes fA: "finite A"  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
333  | 
shows "setsum f (A \<inter> B) = setsum (\<lambda>x. if x \<in> B then f x else 0) A"  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
334  | 
proof-  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
335  | 
from fA have fab: "finite (A \<inter> B)" by auto  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
336  | 
have aba: "A \<inter> B \<subseteq> A" by blast  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
337  | 
let ?g = "\<lambda>x. if x \<in> A\<inter>B then f x else 0"  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
338  | 
from setsum_mono_zero_left[OF fA aba, of ?g]  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
339  | 
show ?thesis by simp  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
340  | 
qed  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
341  | 
|
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
342  | 
lemma setsum_cases:  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
343  | 
assumes fA: "finite A"  | 
| 35577 | 344  | 
shows "setsum (\<lambda>x. if P x then f x else g x) A =  | 
345  | 
         setsum f (A \<inter> {x. P x}) + setsum g (A \<inter> - {x. P x})"
 | 
|
| 42986 | 346  | 
using setsum.If_cases[OF fA] .  | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
347  | 
|
| 
15409
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
348  | 
(*But we can't get rid of finite I. If infinite, although the rhs is 0,  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
349  | 
the lhs need not be, since UNION I A could still be finite.*)  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
350  | 
lemma (in comm_monoid_add) setsum_UN_disjoint:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
351  | 
assumes "finite I" and "ALL i:I. finite (A i)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
352  | 
    and "ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}"
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
353  | 
shows "setsum f (UNION I A) = (\<Sum>i\<in>I. setsum f (A i))"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
354  | 
proof -  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
355  | 
interpret comm_monoid_mult "op +" 0 by (fact comm_monoid_mult)  | 
| 41550 | 356  | 
from assms show ?thesis by (simp add: setsum_def fold_image_UN_disjoint)  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
357  | 
qed  | 
| 15402 | 358  | 
|
| 
15409
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
359  | 
text{*No need to assume that @{term C} is finite.  If infinite, the rhs is
 | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
360  | 
directly 0, and @{term "Union C"} is also infinite, hence the lhs is also 0.*}
 | 
| 15402 | 361  | 
lemma setsum_Union_disjoint:  | 
| 
44937
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
362  | 
  assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A Int B = {}"
 | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
363  | 
shows "setsum f (Union C) = setsum (setsum f) C"  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
364  | 
proof cases  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
365  | 
assume "finite C"  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
366  | 
from setsum_UN_disjoint[OF this assms]  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
367  | 
show ?thesis  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
368  | 
by (simp add: SUP_def)  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
369  | 
qed (force dest: finite_UnionD simp add: setsum_def)  | 
| 15402 | 370  | 
|
| 
15409
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
371  | 
(*But we can't get rid of finite A. If infinite, although the lhs is 0,  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
372  | 
the rhs need not be, since SIGMA A B could still be finite.*)  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
373  | 
lemma (in comm_monoid_add) setsum_Sigma:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
374  | 
assumes "finite A" and "ALL x:A. finite (B x)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
375  | 
shows "(\<Sum>x\<in>A. (\<Sum>y\<in>B x. f x y)) = (\<Sum>(x,y)\<in>(SIGMA x:A. B x). f x y)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
376  | 
proof -  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
377  | 
interpret comm_monoid_mult "op +" 0 by (fact comm_monoid_mult)  | 
| 41550 | 378  | 
from assms show ?thesis by (simp add: setsum_def fold_image_Sigma split_def)  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
379  | 
qed  | 
| 15402 | 380  | 
|
| 
15409
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
381  | 
text{*Here we can eliminate the finiteness assumptions, by cases.*}
 | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
382  | 
lemma setsum_cartesian_product:  | 
| 17189 | 383  | 
"(\<Sum>x\<in>A. (\<Sum>y\<in>B. f x y)) = (\<Sum>(x,y) \<in> A <*> B. f x y)"  | 
| 
15409
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
384  | 
apply (cases "finite A")  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
385  | 
apply (cases "finite B")  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
386  | 
apply (simp add: setsum_Sigma)  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
387  | 
 apply (cases "A={}", simp)
 | 
| 15543 | 388  | 
apply (simp)  | 
| 
15409
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
389  | 
apply (auto simp add: setsum_def  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
390  | 
dest: finite_cartesian_productD1 finite_cartesian_productD2)  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
391  | 
done  | 
| 15402 | 392  | 
|
| 48861 | 393  | 
lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)"  | 
394  | 
by (fact setsum.F_fun_f)  | 
|
| 15402 | 395  | 
|
| 48893 | 396  | 
lemma setsum_Un_zero:  | 
397  | 
"\<lbrakk> finite S; finite T; \<forall>x \<in> S\<inter>T. f x = 0 \<rbrakk> \<Longrightarrow>  | 
|
398  | 
setsum f (S \<union> T) = setsum f S + setsum f T"  | 
|
399  | 
by(fact setsum.F_Un_neutral)  | 
|
400  | 
||
401  | 
lemma setsum_UNION_zero:  | 
|
402  | 
assumes fS: "finite S" and fSS: "\<forall>T \<in> S. finite T"  | 
|
403  | 
and f0: "\<And>T1 T2 x. T1\<in>S \<Longrightarrow> T2\<in>S \<Longrightarrow> T1 \<noteq> T2 \<Longrightarrow> x \<in> T1 \<Longrightarrow> x \<in> T2 \<Longrightarrow> f x = 0"  | 
|
404  | 
shows "setsum f (\<Union>S) = setsum (\<lambda>T. setsum f T) S"  | 
|
405  | 
using fSS f0  | 
|
406  | 
proof(induct rule: finite_induct[OF fS])  | 
|
407  | 
case 1 thus ?case by simp  | 
|
408  | 
next  | 
|
409  | 
case (2 T F)  | 
|
410  | 
then have fTF: "finite T" "\<forall>T\<in>F. finite T" "finite F" and TF: "T \<notin> F"  | 
|
411  | 
and H: "setsum f (\<Union> F) = setsum (setsum f) F" by auto  | 
|
412  | 
from fTF have fUF: "finite (\<Union>F)" by auto  | 
|
413  | 
from "2.prems" TF fTF  | 
|
414  | 
show ?case  | 
|
415  | 
by (auto simp add: H[symmetric] intro: setsum_Un_zero[OF fTF(1) fUF, of f])  | 
|
416  | 
qed  | 
|
417  | 
||
| 15402 | 418  | 
|
419  | 
subsubsection {* Properties in more restricted classes of structures *}
 | 
|
420  | 
||
421  | 
lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a"  | 
|
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
422  | 
apply (case_tac "finite A")  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
423  | 
prefer 2 apply (simp add: setsum_def)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
424  | 
apply (erule rev_mp)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
425  | 
apply (erule finite_induct, auto)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
426  | 
done  | 
| 15402 | 427  | 
|
428  | 
lemma setsum_eq_0_iff [simp]:  | 
|
429  | 
"finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))"  | 
|
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
430  | 
by (induct set: finite) auto  | 
| 15402 | 431  | 
|
| 30859 | 432  | 
lemma setsum_eq_Suc0_iff: "finite A \<Longrightarrow>  | 
433  | 
(setsum f A = Suc 0) = (EX a:A. f a = Suc 0 & (ALL b:A. a\<noteq>b \<longrightarrow> f b = 0))"  | 
|
434  | 
apply(erule finite_induct)  | 
|
435  | 
apply (auto simp add:add_is_1)  | 
|
436  | 
done  | 
|
437  | 
||
438  | 
lemmas setsum_eq_1_iff = setsum_eq_Suc0_iff[simplified One_nat_def[symmetric]]  | 
|
439  | 
||
| 15402 | 440  | 
lemma setsum_Un_nat: "finite A ==> finite B ==>  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
441  | 
(setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"  | 
| 15402 | 442  | 
  -- {* For the natural numbers, we have subtraction. *}
 | 
| 29667 | 443  | 
by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)  | 
| 15402 | 444  | 
|
445  | 
lemma setsum_Un: "finite A ==> finite B ==>  | 
|
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
446  | 
(setsum f (A Un B) :: 'a :: ab_group_add) =  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
447  | 
setsum f A + setsum f B - setsum f (A Int B)"  | 
| 29667 | 448  | 
by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)  | 
| 15402 | 449  | 
|
| 49715 | 450  | 
lemma setsum_Un2:  | 
451  | 
assumes "finite (A \<union> B)"  | 
|
452  | 
shows "setsum f (A \<union> B) = setsum f (A - B) + setsum f (B - A) + setsum f (A \<inter> B)"  | 
|
453  | 
proof -  | 
|
454  | 
have "A \<union> B = A - B \<union> (B - A) \<union> A \<inter> B"  | 
|
455  | 
by auto  | 
|
456  | 
with assms show ?thesis by simp (subst setsum_Un_disjoint, auto)+  | 
|
457  | 
qed  | 
|
458  | 
||
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
459  | 
lemma (in comm_monoid_add) setsum_eq_general_reverses:  | 
| 
30260
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
460  | 
assumes fS: "finite S" and fT: "finite T"  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
461  | 
and kh: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S \<and> h (k y) = y"  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
462  | 
and hk: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> T \<and> k (h x) = x \<and> g (h x) = f x"  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
463  | 
shows "setsum f S = setsum g T"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
464  | 
proof -  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
465  | 
interpret comm_monoid_mult "op +" 0 by (fact comm_monoid_mult)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
466  | 
show ?thesis  | 
| 
30260
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
467  | 
apply (simp add: setsum_def fS fT)  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
468  | 
apply (rule fold_image_eq_general_inverses)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
469  | 
apply (rule fS)  | 
| 
30260
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
470  | 
apply (erule kh)  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
471  | 
apply (erule hk)  | 
| 
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
472  | 
done  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
473  | 
qed  | 
| 
30260
 
be39acd3ac85
Added general theorems for fold_image, setsum and set_prod
 
chaieb 
parents: 
29966 
diff
changeset
 | 
474  | 
|
| 15402 | 475  | 
lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) =
 | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
476  | 
(if a:A then setsum f A - f a else setsum f A)"  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
477  | 
apply (case_tac "finite A")  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
478  | 
prefer 2 apply (simp add: setsum_def)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
479  | 
apply (erule finite_induct)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
480  | 
apply (auto simp add: insert_Diff_if)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
481  | 
apply (drule_tac a = a in mk_disjoint_insert, auto)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
482  | 
done  | 
| 15402 | 483  | 
|
484  | 
lemma setsum_diff1: "finite A \<Longrightarrow>  | 
|
485  | 
  (setsum f (A - {a}) :: ('a::ab_group_add)) =
 | 
|
486  | 
(if a:A then setsum f A - f a else setsum f A)"  | 
|
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
487  | 
by (erule finite_induct) (auto simp add: insert_Diff_if)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
488  | 
|
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
489  | 
lemma setsum_diff1'[rule_format]:  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
490  | 
  "finite A \<Longrightarrow> a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x)"
 | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
491  | 
apply (erule finite_induct[where F=A and P="% A. (a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x))"])
 | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
492  | 
apply (auto simp add: insert_Diff_if add_ac)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
493  | 
done  | 
| 
15552
 
8ab8e425410b
added setsum_diff1' which holds in more general cases than setsum_diff1
 
obua 
parents: 
15543 
diff
changeset
 | 
494  | 
|
| 31438 | 495  | 
lemma setsum_diff1_ring: assumes "finite A" "a \<in> A"  | 
496  | 
  shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)"
 | 
|
497  | 
unfolding setsum_diff1'[OF assms] by auto  | 
|
498  | 
||
| 15402 | 499  | 
(* By Jeremy Siek: *)  | 
500  | 
||
501  | 
lemma setsum_diff_nat:  | 
|
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
502  | 
assumes "finite B" and "B \<subseteq> A"  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
503  | 
shows "(setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)"  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
504  | 
using assms  | 
| 19535 | 505  | 
proof induct  | 
| 15402 | 506  | 
  show "setsum f (A - {}) = (setsum f A) - (setsum f {})" by simp
 | 
507  | 
next  | 
|
508  | 
fix F x assume finF: "finite F" and xnotinF: "x \<notin> F"  | 
|
509  | 
and xFinA: "insert x F \<subseteq> A"  | 
|
510  | 
and IH: "F \<subseteq> A \<Longrightarrow> setsum f (A - F) = setsum f A - setsum f F"  | 
|
511  | 
from xnotinF xFinA have xinAF: "x \<in> (A - F)" by simp  | 
|
512  | 
  from xinAF have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x"
 | 
|
513  | 
by (simp add: setsum_diff1_nat)  | 
|
514  | 
from xFinA have "F \<subseteq> A" by simp  | 
|
515  | 
with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp  | 
|
516  | 
  with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x"
 | 
|
517  | 
by simp  | 
|
518  | 
  from xnotinF have "A - insert x F = (A - F) - {x}" by auto
 | 
|
519  | 
with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x"  | 
|
520  | 
by simp  | 
|
521  | 
from finF xnotinF have "setsum f (insert x F) = setsum f F + f x" by simp  | 
|
522  | 
with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)"  | 
|
523  | 
by simp  | 
|
524  | 
thus "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp  | 
|
525  | 
qed  | 
|
526  | 
||
527  | 
lemma setsum_diff:  | 
|
528  | 
assumes le: "finite A" "B \<subseteq> A"  | 
|
529  | 
  shows "setsum f (A - B) = setsum f A - ((setsum f B)::('a::ab_group_add))"
 | 
|
530  | 
proof -  | 
|
531  | 
from le have finiteB: "finite B" using finite_subset by auto  | 
|
532  | 
show ?thesis using finiteB le  | 
|
| 21575 | 533  | 
proof induct  | 
| 19535 | 534  | 
case empty  | 
535  | 
thus ?case by auto  | 
|
536  | 
next  | 
|
537  | 
case (insert x F)  | 
|
538  | 
thus ?case using le finiteB  | 
|
539  | 
by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb)  | 
|
| 15402 | 540  | 
qed  | 
| 19535 | 541  | 
qed  | 
| 15402 | 542  | 
|
543  | 
lemma setsum_mono:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
544  | 
  assumes le: "\<And>i. i\<in>K \<Longrightarrow> f (i::'a) \<le> ((g i)::('b::{comm_monoid_add, ordered_ab_semigroup_add}))"
 | 
| 15402 | 545  | 
shows "(\<Sum>i\<in>K. f i) \<le> (\<Sum>i\<in>K. g i)"  | 
546  | 
proof (cases "finite K")  | 
|
547  | 
case True  | 
|
548  | 
thus ?thesis using le  | 
|
| 19535 | 549  | 
proof induct  | 
| 15402 | 550  | 
case empty  | 
551  | 
thus ?case by simp  | 
|
552  | 
next  | 
|
553  | 
case insert  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44845 
diff
changeset
 | 
554  | 
thus ?case using add_mono by fastforce  | 
| 15402 | 555  | 
qed  | 
556  | 
next  | 
|
557  | 
case False  | 
|
558  | 
thus ?thesis  | 
|
559  | 
by (simp add: setsum_def)  | 
|
560  | 
qed  | 
|
561  | 
||
| 15554 | 562  | 
lemma setsum_strict_mono:  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
563  | 
  fixes f :: "'a \<Rightarrow> 'b::{ordered_cancel_ab_semigroup_add,comm_monoid_add}"
 | 
| 19535 | 564  | 
  assumes "finite A"  "A \<noteq> {}"
 | 
565  | 
and "!!x. x:A \<Longrightarrow> f x < g x"  | 
|
566  | 
shows "setsum f A < setsum g A"  | 
|
| 41550 | 567  | 
using assms  | 
| 15554 | 568  | 
proof (induct rule: finite_ne_induct)  | 
569  | 
case singleton thus ?case by simp  | 
|
570  | 
next  | 
|
571  | 
case insert thus ?case by (auto simp: add_strict_mono)  | 
|
572  | 
qed  | 
|
573  | 
||
| 46699 | 574  | 
lemma setsum_strict_mono_ex1:  | 
575  | 
fixes f :: "'a \<Rightarrow> 'b::{comm_monoid_add, ordered_cancel_ab_semigroup_add}"
 | 
|
576  | 
assumes "finite A" and "ALL x:A. f x \<le> g x" and "EX a:A. f a < g a"  | 
|
577  | 
shows "setsum f A < setsum g A"  | 
|
578  | 
proof-  | 
|
579  | 
from assms(3) obtain a where a: "a:A" "f a < g a" by blast  | 
|
580  | 
  have "setsum f A = setsum f ((A-{a}) \<union> {a})"
 | 
|
581  | 
by(simp add:insert_absorb[OF `a:A`])  | 
|
582  | 
  also have "\<dots> = setsum f (A-{a}) + setsum f {a}"
 | 
|
583  | 
using `finite A` by(subst setsum_Un_disjoint) auto  | 
|
584  | 
  also have "setsum f (A-{a}) \<le> setsum g (A-{a})"
 | 
|
585  | 
by(rule setsum_mono)(simp add: assms(2))  | 
|
586  | 
  also have "setsum f {a} < setsum g {a}" using a by simp
 | 
|
587  | 
  also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \<union> {a})"
 | 
|
588  | 
using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto  | 
|
589  | 
also have "\<dots> = setsum g A" by(simp add:insert_absorb[OF `a:A`])  | 
|
590  | 
finally show ?thesis by (metis add_right_mono add_strict_left_mono)  | 
|
591  | 
qed  | 
|
592  | 
||
| 15535 | 593  | 
lemma setsum_negf:  | 
| 19535 | 594  | 
"setsum (%x. - (f x)::'a::ab_group_add) A = - setsum f A"  | 
| 15535 | 595  | 
proof (cases "finite A")  | 
| 22262 | 596  | 
case True thus ?thesis by (induct set: finite) auto  | 
| 15535 | 597  | 
next  | 
598  | 
case False thus ?thesis by (simp add: setsum_def)  | 
|
599  | 
qed  | 
|
| 15402 | 600  | 
|
| 15535 | 601  | 
lemma setsum_subtractf:  | 
| 19535 | 602  | 
"setsum (%x. ((f x)::'a::ab_group_add) - g x) A =  | 
603  | 
setsum f A - setsum g A"  | 
|
| 15535 | 604  | 
proof (cases "finite A")  | 
605  | 
case True thus ?thesis by (simp add: diff_minus setsum_addf setsum_negf)  | 
|
606  | 
next  | 
|
607  | 
case False thus ?thesis by (simp add: setsum_def)  | 
|
608  | 
qed  | 
|
| 15402 | 609  | 
|
| 15535 | 610  | 
lemma setsum_nonneg:  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
611  | 
  assumes nn: "\<forall>x\<in>A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \<le> f x"
 | 
| 19535 | 612  | 
shows "0 \<le> setsum f A"  | 
| 15535 | 613  | 
proof (cases "finite A")  | 
614  | 
case True thus ?thesis using nn  | 
|
| 21575 | 615  | 
proof induct  | 
| 19535 | 616  | 
case empty then show ?case by simp  | 
617  | 
next  | 
|
618  | 
case (insert x F)  | 
|
619  | 
then have "0 + 0 \<le> f x + setsum f F" by (blast intro: add_mono)  | 
|
620  | 
with insert show ?case by simp  | 
|
621  | 
qed  | 
|
| 15535 | 622  | 
next  | 
623  | 
case False thus ?thesis by (simp add: setsum_def)  | 
|
624  | 
qed  | 
|
| 15402 | 625  | 
|
| 15535 | 626  | 
lemma setsum_nonpos:  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
627  | 
  assumes np: "\<forall>x\<in>A. f x \<le> (0::'a::{ordered_ab_semigroup_add,comm_monoid_add})"
 | 
| 19535 | 628  | 
shows "setsum f A \<le> 0"  | 
| 15535 | 629  | 
proof (cases "finite A")  | 
630  | 
case True thus ?thesis using np  | 
|
| 21575 | 631  | 
proof induct  | 
| 19535 | 632  | 
case empty then show ?case by simp  | 
633  | 
next  | 
|
634  | 
case (insert x F)  | 
|
635  | 
then have "f x + setsum f F \<le> 0 + 0" by (blast intro: add_mono)  | 
|
636  | 
with insert show ?case by simp  | 
|
637  | 
qed  | 
|
| 15535 | 638  | 
next  | 
639  | 
case False thus ?thesis by (simp add: setsum_def)  | 
|
640  | 
qed  | 
|
| 15402 | 641  | 
|
| 
36622
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
642  | 
lemma setsum_nonneg_leq_bound:  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
643  | 
  fixes f :: "'a \<Rightarrow> 'b::{ordered_ab_group_add}"
 | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
644  | 
assumes "finite s" "\<And>i. i \<in> s \<Longrightarrow> f i \<ge> 0" "(\<Sum>i \<in> s. f i) = B" "i \<in> s"  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
645  | 
shows "f i \<le> B"  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
646  | 
proof -  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
647  | 
  have "0 \<le> (\<Sum> i \<in> s - {i}. f i)" and "0 \<le> f i"
 | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
648  | 
using assms by (auto intro!: setsum_nonneg)  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
649  | 
moreover  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
650  | 
  have "(\<Sum> i \<in> s - {i}. f i) + f i = B"
 | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
651  | 
using assms by (simp add: setsum_diff1)  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
652  | 
ultimately show ?thesis by auto  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
653  | 
qed  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
654  | 
|
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
655  | 
lemma setsum_nonneg_0:  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
656  | 
  fixes f :: "'a \<Rightarrow> 'b::{ordered_ab_group_add}"
 | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
657  | 
assumes "finite s" and pos: "\<And> i. i \<in> s \<Longrightarrow> f i \<ge> 0"  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
658  | 
and "(\<Sum> i \<in> s. f i) = 0" and i: "i \<in> s"  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
659  | 
shows "f i = 0"  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
660  | 
using setsum_nonneg_leq_bound[OF assms] pos[OF i] by auto  | 
| 
 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 
hoelzl 
parents: 
36409 
diff
changeset
 | 
661  | 
|
| 15539 | 662  | 
lemma setsum_mono2:  | 
| 36303 | 663  | 
fixes f :: "'a \<Rightarrow> 'b :: ordered_comm_monoid_add"  | 
| 15539 | 664  | 
assumes fin: "finite B" and sub: "A \<subseteq> B" and nn: "\<And>b. b \<in> B-A \<Longrightarrow> 0 \<le> f b"  | 
665  | 
shows "setsum f A \<le> setsum f B"  | 
|
666  | 
proof -  | 
|
667  | 
have "setsum f A \<le> setsum f A + setsum f (B-A)"  | 
|
668  | 
by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def)  | 
|
669  | 
also have "\<dots> = setsum f (A \<union> (B-A))" using fin finite_subset[OF sub fin]  | 
|
670  | 
by (simp add:setsum_Un_disjoint del:Un_Diff_cancel)  | 
|
671  | 
also have "A \<union> (B-A) = B" using sub by blast  | 
|
672  | 
finally show ?thesis .  | 
|
673  | 
qed  | 
|
| 15542 | 674  | 
|
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
675  | 
lemma setsum_mono3: "finite B ==> A <= B ==>  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
676  | 
ALL x: B - A.  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
677  | 
      0 <= ((f x)::'a::{comm_monoid_add,ordered_ab_semigroup_add}) ==>
 | 
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
678  | 
setsum f A <= setsum f B"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
679  | 
apply (subgoal_tac "setsum f B = setsum f A + setsum f (B - A)")  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
680  | 
apply (erule ssubst)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
681  | 
apply (subgoal_tac "setsum f A + 0 <= setsum f A + setsum f (B - A)")  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
682  | 
apply simp  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
683  | 
apply (rule add_left_mono)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
684  | 
apply (erule setsum_nonneg)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
685  | 
apply (subst setsum_Un_disjoint [THEN sym])  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
686  | 
apply (erule finite_subset, assumption)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
687  | 
apply (rule finite_subset)  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
688  | 
prefer 2  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
689  | 
apply assumption  | 
| 
32698
 
be4b248616c0
inf/sup_absorb are no default simp rules any longer
 
haftmann 
parents: 
32697 
diff
changeset
 | 
690  | 
apply (auto simp add: sup_absorb2)  | 
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
691  | 
done  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
692  | 
|
| 19279 | 693  | 
lemma setsum_right_distrib:  | 
| 
22934
 
64ecb3d6790a
generalize setsum lemmas from semiring_0_cancel to semiring_0
 
huffman 
parents: 
22917 
diff
changeset
 | 
694  | 
  fixes f :: "'a => ('b::semiring_0)"
 | 
| 15402 | 695  | 
shows "r * setsum f A = setsum (%n. r * f n) A"  | 
696  | 
proof (cases "finite A")  | 
|
697  | 
case True  | 
|
698  | 
thus ?thesis  | 
|
| 21575 | 699  | 
proof induct  | 
| 15402 | 700  | 
case empty thus ?case by simp  | 
701  | 
next  | 
|
702  | 
case (insert x A) thus ?case by (simp add: right_distrib)  | 
|
703  | 
qed  | 
|
704  | 
next  | 
|
705  | 
case False thus ?thesis by (simp add: setsum_def)  | 
|
706  | 
qed  | 
|
707  | 
||
| 
17149
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
708  | 
lemma setsum_left_distrib:  | 
| 
22934
 
64ecb3d6790a
generalize setsum lemmas from semiring_0_cancel to semiring_0
 
huffman 
parents: 
22917 
diff
changeset
 | 
709  | 
"setsum f A * (r::'a::semiring_0) = (\<Sum>n\<in>A. f n * r)"  | 
| 
17149
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
710  | 
proof (cases "finite A")  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
711  | 
case True  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
712  | 
then show ?thesis  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
713  | 
proof induct  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
714  | 
case empty thus ?case by simp  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
715  | 
next  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
716  | 
case (insert x A) thus ?case by (simp add: left_distrib)  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
717  | 
qed  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
718  | 
next  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
719  | 
case False thus ?thesis by (simp add: setsum_def)  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
720  | 
qed  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
721  | 
|
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
722  | 
lemma setsum_divide_distrib:  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
723  | 
"setsum f A / (r::'a::field) = (\<Sum>n\<in>A. f n / r)"  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
724  | 
proof (cases "finite A")  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
725  | 
case True  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
726  | 
then show ?thesis  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
727  | 
proof induct  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
728  | 
case empty thus ?case by simp  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
729  | 
next  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
730  | 
case (insert x A) thus ?case by (simp add: add_divide_distrib)  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
731  | 
qed  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
732  | 
next  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
733  | 
case False thus ?thesis by (simp add: setsum_def)  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
734  | 
qed  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
735  | 
|
| 15535 | 736  | 
lemma setsum_abs[iff]:  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
737  | 
  fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
 | 
| 15402 | 738  | 
shows "abs (setsum f A) \<le> setsum (%i. abs(f i)) A"  | 
| 15535 | 739  | 
proof (cases "finite A")  | 
740  | 
case True  | 
|
741  | 
thus ?thesis  | 
|
| 21575 | 742  | 
proof induct  | 
| 15535 | 743  | 
case empty thus ?case by simp  | 
744  | 
next  | 
|
745  | 
case (insert x A)  | 
|
746  | 
thus ?case by (auto intro: abs_triangle_ineq order_trans)  | 
|
747  | 
qed  | 
|
| 15402 | 748  | 
next  | 
| 15535 | 749  | 
case False thus ?thesis by (simp add: setsum_def)  | 
| 15402 | 750  | 
qed  | 
751  | 
||
| 15535 | 752  | 
lemma setsum_abs_ge_zero[iff]:  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
753  | 
  fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
 | 
| 15402 | 754  | 
shows "0 \<le> setsum (%i. abs(f i)) A"  | 
| 15535 | 755  | 
proof (cases "finite A")  | 
756  | 
case True  | 
|
757  | 
thus ?thesis  | 
|
| 21575 | 758  | 
proof induct  | 
| 15535 | 759  | 
case empty thus ?case by simp  | 
760  | 
next  | 
|
| 
36977
 
71c8973a604b
declare add_nonneg_nonneg [simp]; remove now-redundant lemmas realpow_two_le_order(2)
 
huffman 
parents: 
36635 
diff
changeset
 | 
761  | 
case (insert x A) thus ?case by auto  | 
| 15535 | 762  | 
qed  | 
| 15402 | 763  | 
next  | 
| 15535 | 764  | 
case False thus ?thesis by (simp add: setsum_def)  | 
| 15402 | 765  | 
qed  | 
766  | 
||
| 15539 | 767  | 
lemma abs_setsum_abs[simp]:  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
768  | 
  fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
 | 
| 15539 | 769  | 
shows "abs (\<Sum>a\<in>A. abs(f a)) = (\<Sum>a\<in>A. abs(f a))"  | 
770  | 
proof (cases "finite A")  | 
|
771  | 
case True  | 
|
772  | 
thus ?thesis  | 
|
| 21575 | 773  | 
proof induct  | 
| 15539 | 774  | 
case empty thus ?case by simp  | 
775  | 
next  | 
|
776  | 
case (insert a A)  | 
|
777  | 
hence "\<bar>\<Sum>a\<in>insert a A. \<bar>f a\<bar>\<bar> = \<bar>\<bar>f a\<bar> + (\<Sum>a\<in>A. \<bar>f a\<bar>)\<bar>" by simp  | 
|
778  | 
also have "\<dots> = \<bar>\<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>\<bar>" using insert by simp  | 
|
| 
16775
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
779  | 
also have "\<dots> = \<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>"  | 
| 
 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 
avigad 
parents: 
16760 
diff
changeset
 | 
780  | 
by (simp del: abs_of_nonneg)  | 
| 15539 | 781  | 
also have "\<dots> = (\<Sum>a\<in>insert a A. \<bar>f a\<bar>)" using insert by simp  | 
782  | 
finally show ?case .  | 
|
783  | 
qed  | 
|
784  | 
next  | 
|
785  | 
case False thus ?thesis by (simp add: setsum_def)  | 
|
786  | 
qed  | 
|
787  | 
||
| 31080 | 788  | 
lemma setsum_Plus:  | 
789  | 
fixes A :: "'a set" and B :: "'b set"  | 
|
790  | 
assumes fin: "finite A" "finite B"  | 
|
791  | 
shows "setsum f (A <+> B) = setsum (f \<circ> Inl) A + setsum (f \<circ> Inr) B"  | 
|
792  | 
proof -  | 
|
793  | 
have "A <+> B = Inl ` A \<union> Inr ` B" by auto  | 
|
794  | 
  moreover from fin have "finite (Inl ` A :: ('a + 'b) set)" "finite (Inr ` B :: ('a + 'b) set)"
 | 
|
| 
40786
 
0a54cfc9add3
gave more standard finite set rules simp and intro attribute
 
nipkow 
parents: 
39302 
diff
changeset
 | 
795  | 
by auto  | 
| 31080 | 796  | 
  moreover have "Inl ` A \<inter> Inr ` B = ({} :: ('a + 'b) set)" by auto
 | 
797  | 
moreover have "inj_on (Inl :: 'a \<Rightarrow> 'a + 'b) A" "inj_on (Inr :: 'b \<Rightarrow> 'a + 'b) B" by(auto intro: inj_onI)  | 
|
798  | 
ultimately show ?thesis using fin by(simp add: setsum_Un_disjoint setsum_reindex)  | 
|
799  | 
qed  | 
|
800  | 
||
801  | 
||
| 
17149
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
802  | 
text {* Commuting outer and inner summation *}
 | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
803  | 
|
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
804  | 
lemma setsum_commute:  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
805  | 
"(\<Sum>i\<in>A. \<Sum>j\<in>B. f i j) = (\<Sum>j\<in>B. \<Sum>i\<in>A. f i j)"  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
806  | 
proof (simp add: setsum_cartesian_product)  | 
| 17189 | 807  | 
have "(\<Sum>(x,y) \<in> A <*> B. f x y) =  | 
808  | 
(\<Sum>(y,x) \<in> (%(i, j). (j, i)) ` (A \<times> B). f x y)"  | 
|
| 
17149
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
809  | 
(is "?s = _")  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
810  | 
apply (simp add: setsum_reindex [where f = "%(i, j). (j, i)"] swap_inj_on)  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
811  | 
apply (simp add: split_def)  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
812  | 
done  | 
| 17189 | 813  | 
also have "... = (\<Sum>(y,x)\<in>B \<times> A. f x y)"  | 
| 
17149
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
814  | 
(is "_ = ?t")  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
815  | 
apply (simp add: swap_product)  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
816  | 
done  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
817  | 
finally show "?s = ?t" .  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
818  | 
qed  | 
| 
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
819  | 
|
| 19279 | 820  | 
lemma setsum_product:  | 
| 
22934
 
64ecb3d6790a
generalize setsum lemmas from semiring_0_cancel to semiring_0
 
huffman 
parents: 
22917 
diff
changeset
 | 
821  | 
  fixes f :: "'a => ('b::semiring_0)"
 | 
| 19279 | 822  | 
shows "setsum f A * setsum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"  | 
823  | 
by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum_commute)  | 
|
824  | 
||
| 34223 | 825  | 
lemma setsum_mult_setsum_if_inj:  | 
826  | 
fixes f :: "'a => ('b::semiring_0)"
 | 
|
827  | 
shows "inj_on (%(a,b). f a * g b) (A \<times> B) ==>  | 
|
828  | 
  setsum f A * setsum g B = setsum id {f a * g b|a b. a:A & b:B}"
 | 
|
829  | 
by(auto simp: setsum_product setsum_cartesian_product  | 
|
830  | 
intro!: setsum_reindex_cong[symmetric])  | 
|
831  | 
||
| 
35722
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
832  | 
lemma setsum_constant [simp]: "(\<Sum>x \<in> A. y) = of_nat(card A) * y"  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
833  | 
apply (cases "finite A")  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
834  | 
apply (erule finite_induct)  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
835  | 
apply (auto simp add: algebra_simps)  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
836  | 
done  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
837  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
838  | 
lemma setsum_bounded:  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
839  | 
  assumes le: "\<And>i. i\<in>A \<Longrightarrow> f i \<le> (K::'a::{semiring_1, ordered_ab_semigroup_add})"
 | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
840  | 
shows "setsum f A \<le> of_nat(card A) * K"  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
841  | 
proof (cases "finite A")  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
842  | 
case True  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
843  | 
thus ?thesis using le setsum_mono[where K=A and g = "%x. K"] by simp  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
844  | 
next  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
845  | 
case False thus ?thesis by (simp add: setsum_def)  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
846  | 
qed  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
847  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
848  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
849  | 
subsubsection {* Cardinality as special case of @{const setsum} *}
 | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
850  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
851  | 
lemma card_eq_setsum:  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
852  | 
"card A = setsum (\<lambda>x. 1) A"  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
853  | 
by (simp only: card_def setsum_def)  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
854  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
855  | 
lemma card_UN_disjoint:  | 
| 46629 | 856  | 
assumes "finite I" and "\<forall>i\<in>I. finite (A i)"  | 
857  | 
    and "\<forall>i\<in>I. \<forall>j\<in>I. i \<noteq> j \<longrightarrow> A i \<inter> A j = {}"
 | 
|
858  | 
shows "card (UNION I A) = (\<Sum>i\<in>I. card(A i))"  | 
|
859  | 
proof -  | 
|
860  | 
have "(\<Sum>i\<in>I. card (A i)) = (\<Sum>i\<in>I. \<Sum>x\<in>A i. 1)" by simp  | 
|
861  | 
with assms show ?thesis by (simp add: card_eq_setsum setsum_UN_disjoint del: setsum_constant)  | 
|
862  | 
qed  | 
|
| 
35722
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
863  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
864  | 
lemma card_Union_disjoint:  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
865  | 
"finite C ==> (ALL A:C. finite A) ==>  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
866  | 
   (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {})
 | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
867  | 
==> card (Union C) = setsum card C"  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
868  | 
apply (frule card_UN_disjoint [of C id])  | 
| 
44937
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
869  | 
apply (simp_all add: SUP_def id_def)  | 
| 
35722
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
870  | 
done  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
871  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
872  | 
text{*The image of a finite set can be expressed using @{term fold_image}.*}
 | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
873  | 
lemma image_eq_fold_image:  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
874  | 
  "finite A ==> f ` A = fold_image (op Un) (%x. {f x}) {} A"
 | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
875  | 
proof (induct rule: finite_induct)  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
876  | 
case empty then show ?case by simp  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
877  | 
next  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
878  | 
interpret ab_semigroup_mult "op Un"  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
879  | 
proof qed auto  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
880  | 
case insert  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
881  | 
then show ?case by simp  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
882  | 
qed  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
883  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
884  | 
subsubsection {* Cardinality of products *}
 | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
885  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
886  | 
lemma card_SigmaI [simp]:  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
887  | 
"\<lbrakk> finite A; ALL a:A. finite (B a) \<rbrakk>  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
888  | 
\<Longrightarrow> card (SIGMA x: A. B x) = (\<Sum>a\<in>A. card (B a))"  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
889  | 
by(simp add: card_eq_setsum setsum_Sigma del:setsum_constant)  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
890  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
891  | 
(*  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
892  | 
lemma SigmaI_insert: "y \<notin> A ==>  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
893  | 
  (SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \<union> (SIGMA x: A. B x))"
 | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
894  | 
by auto  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
895  | 
*)  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
896  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
897  | 
lemma card_cartesian_product: "card (A <*> B) = card(A) * card(B)"  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
898  | 
by (cases "finite A \<and> finite B")  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
899  | 
(auto simp add: card_eq_0_iff dest: finite_cartesian_productD1 finite_cartesian_productD2)  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
900  | 
|
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
901  | 
lemma card_cartesian_product_singleton:  "card({x} <*> A) = card(A)"
 | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
902  | 
by (simp add: card_cartesian_product)  | 
| 
 
69419a09a7ff
moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
 
haftmann 
parents: 
35719 
diff
changeset
 | 
903  | 
|
| 
17149
 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 
ballarin 
parents: 
17085 
diff
changeset
 | 
904  | 
|
| 15402 | 905  | 
subsection {* Generalized product over a set *}
 | 
906  | 
||
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
907  | 
definition (in comm_monoid_mult) setprod :: "('b \<Rightarrow> 'a) => 'b set => 'a" where
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
908  | 
"setprod f A = (if finite A then fold_image (op *) f 1 A else 1)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
909  | 
|
| 
35938
 
93faaa15c3d5
sublocale comm_monoid_add < setprod --> sublocale comm_monoid_mult < setprod
 
huffman 
parents: 
35831 
diff
changeset
 | 
910  | 
sublocale comm_monoid_mult < setprod!: comm_monoid_big "op *" 1 setprod proof  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
911  | 
qed (fact setprod_def)  | 
| 15402 | 912  | 
|
| 19535 | 913  | 
abbreviation  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21249 
diff
changeset
 | 
914  | 
  Setprod  ("\<Prod>_" [1000] 999) where
 | 
| 19535 | 915  | 
"\<Prod>A == setprod (%x. x) A"  | 
916  | 
||
| 15402 | 917  | 
syntax  | 
| 17189 | 918  | 
  "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"  ("(3PROD _:_. _)" [0, 51, 10] 10)
 | 
| 15402 | 919  | 
syntax (xsymbols)  | 
| 17189 | 920  | 
  "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"  ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
 | 
| 15402 | 921  | 
syntax (HTML output)  | 
| 17189 | 922  | 
  "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"  ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
 | 
| 16550 | 923  | 
|
924  | 
translations -- {* Beware of argument permutation! *}
 | 
|
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
925  | 
"PROD i:A. b" == "CONST setprod (%i. b) A"  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
926  | 
"\<Prod>i\<in>A. b" == "CONST setprod (%i. b) A"  | 
| 16550 | 927  | 
|
928  | 
text{* Instead of @{term"\<Prod>x\<in>{x. P}. e"} we introduce the shorter
 | 
|
929  | 
 @{text"\<Prod>x|P. e"}. *}
 | 
|
930  | 
||
931  | 
syntax  | 
|
| 17189 | 932  | 
  "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3PROD _ |/ _./ _)" [0,0,10] 10)
 | 
| 16550 | 933  | 
syntax (xsymbols)  | 
| 17189 | 934  | 
  "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
 | 
| 16550 | 935  | 
syntax (HTML output)  | 
| 17189 | 936  | 
  "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
 | 
| 16550 | 937  | 
|
| 15402 | 938  | 
translations  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
939  | 
  "PROD x|P. t" => "CONST setprod (%x. t) {x. P}"
 | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
940  | 
  "\<Prod>x|P. t" => "CONST setprod (%x. t) {x. P}"
 | 
| 16550 | 941  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
942  | 
lemma setprod_empty: "setprod f {} = 1"
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
943  | 
by (fact setprod.empty)  | 
| 15402 | 944  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
945  | 
lemma setprod_insert: "[| finite A; a \<notin> A |] ==>  | 
| 15402 | 946  | 
setprod f (insert a A) = f a * setprod f A"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
947  | 
by (fact setprod.insert)  | 
| 15402 | 948  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
949  | 
lemma setprod_infinite: "~ finite A ==> setprod f A = 1"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
950  | 
by (fact setprod.infinite)  | 
| 
15409
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
951  | 
|
| 15402 | 952  | 
lemma setprod_reindex:  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
953  | 
"inj_on f B ==> setprod h (f ` B) = setprod (h \<circ> f) B"  | 
| 48849 | 954  | 
by(auto simp: setprod_def fold_image_reindex o_def dest!:finite_imageD)  | 
| 15402 | 955  | 
|
956  | 
lemma setprod_reindex_id: "inj_on f B ==> setprod f B = setprod id (f ` B)"  | 
|
957  | 
by (auto simp add: setprod_reindex)  | 
|
958  | 
||
959  | 
lemma setprod_cong:  | 
|
960  | 
"A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B"  | 
|
| 48849 | 961  | 
by(fact setprod.F_cong)  | 
| 15402 | 962  | 
|
| 48849 | 963  | 
lemma strong_setprod_cong:  | 
| 
16632
 
ad2895beef79
Added strong_setsum_cong and strong_setprod_cong.
 
berghofe 
parents: 
16550 
diff
changeset
 | 
964  | 
"A = B ==> (!!x. x:B =simp=> f x = g x) ==> setprod f A = setprod g B"  | 
| 48849 | 965  | 
by(fact setprod.strong_F_cong)  | 
| 
16632
 
ad2895beef79
Added strong_setsum_cong and strong_setprod_cong.
 
berghofe 
parents: 
16550 
diff
changeset
 | 
966  | 
|
| 15402 | 967  | 
lemma setprod_reindex_cong: "inj_on f A ==>  | 
968  | 
B = f ` A ==> g = h \<circ> f ==> setprod h B = setprod g A"  | 
|
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
969  | 
by (frule setprod_reindex, simp)  | 
| 15402 | 970  | 
|
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
971  | 
lemma strong_setprod_reindex_cong: assumes i: "inj_on f A"  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
972  | 
and B: "B = f ` A" and eq: "\<And>x. x \<in> A \<Longrightarrow> g x = (h \<circ> f) x"  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
973  | 
shows "setprod h B = setprod g A"  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
974  | 
proof-  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
975  | 
have "setprod h B = setprod (h o f) A"  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
976  | 
by (simp add: B setprod_reindex[OF i, of h])  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
977  | 
then show ?thesis apply simp  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
978  | 
apply (rule setprod_cong)  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
979  | 
apply simp  | 
| 
30837
 
3d4832d9f7e4
added strong_setprod_cong[cong] (in analogy with setsum)
 
nipkow 
parents: 
30729 
diff
changeset
 | 
980  | 
by (simp add: eq)  | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
981  | 
qed  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
982  | 
|
| 48893 | 983  | 
lemma setprod_Un_one: "\<lbrakk> finite S; finite T; \<forall>x \<in> S\<inter>T. f x = 1 \<rbrakk>  | 
984  | 
\<Longrightarrow> setprod f (S \<union> T) = setprod f S * setprod f T"  | 
|
985  | 
by(fact setprod.F_Un_neutral)  | 
|
| 15402 | 986  | 
|
| 48821 | 987  | 
lemmas setprod_1 = setprod.F_neutral  | 
988  | 
lemmas setprod_1' = setprod.F_neutral'  | 
|
| 15402 | 989  | 
|
990  | 
||
991  | 
lemma setprod_Un_Int: "finite A ==> finite B  | 
|
992  | 
==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B"  | 
|
| 48849 | 993  | 
by (fact setprod.union_inter)  | 
| 15402 | 994  | 
|
995  | 
lemma setprod_Un_disjoint: "finite A ==> finite B  | 
|
996  | 
  ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
 | 
|
| 48849 | 997  | 
by (fact setprod.union_disjoint)  | 
998  | 
||
999  | 
lemma setprod_subset_diff: "\<lbrakk> B \<subseteq> A; finite A \<rbrakk> \<Longrightarrow>  | 
|
1000  | 
setprod f A = setprod f (A - B) * setprod f B"  | 
|
1001  | 
by(fact setprod.F_subset_diff)  | 
|
| 15402 | 1002  | 
|
| 48849 | 1003  | 
lemma setprod_mono_one_left:  | 
1004  | 
"\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. f i = 1 \<rbrakk> \<Longrightarrow> setprod f S = setprod f T"  | 
|
1005  | 
by(fact setprod.F_mono_neutral_left)  | 
|
| 
30837
 
3d4832d9f7e4
added strong_setprod_cong[cong] (in analogy with setsum)
 
nipkow 
parents: 
30729 
diff
changeset
 | 
1006  | 
|
| 48849 | 1007  | 
lemmas setprod_mono_one_right = setprod.F_mono_neutral_right  | 
| 
30837
 
3d4832d9f7e4
added strong_setprod_cong[cong] (in analogy with setsum)
 
nipkow 
parents: 
30729 
diff
changeset
 | 
1008  | 
|
| 48849 | 1009  | 
lemma setprod_mono_one_cong_left:  | 
1010  | 
"\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 1; \<And>x. x \<in> S \<Longrightarrow> f x = g x \<rbrakk>  | 
|
1011  | 
\<Longrightarrow> setprod f S = setprod g T"  | 
|
1012  | 
by(fact setprod.F_mono_neutral_cong_left)  | 
|
1013  | 
||
1014  | 
lemmas setprod_mono_one_cong_right = setprod.F_mono_neutral_cong_right  | 
|
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1015  | 
|
| 48849 | 1016  | 
lemma setprod_delta: "finite S \<Longrightarrow>  | 
1017  | 
setprod (\<lambda>k. if k=a then b k else 1) S = (if a \<in> S then b a else 1)"  | 
|
1018  | 
by(fact setprod.F_delta)  | 
|
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1019  | 
|
| 48849 | 1020  | 
lemma setprod_delta': "finite S \<Longrightarrow>  | 
1021  | 
setprod (\<lambda>k. if a = k then b k else 1) S = (if a\<in> S then b a else 1)"  | 
|
1022  | 
by(fact setprod.F_delta')  | 
|
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1023  | 
|
| 15402 | 1024  | 
lemma setprod_UN_disjoint:  | 
1025  | 
"finite I ==> (ALL i:I. finite (A i)) ==>  | 
|
1026  | 
        (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
 | 
|
1027  | 
setprod f (UNION I A) = setprod (%i. setprod f (A i)) I"  | 
|
| 41550 | 1028  | 
by (simp add: setprod_def fold_image_UN_disjoint)  | 
| 15402 | 1029  | 
|
1030  | 
lemma setprod_Union_disjoint:  | 
|
| 
44937
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
1031  | 
  assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A Int B = {}" 
 | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
1032  | 
shows "setprod f (Union C) = setprod (setprod f) C"  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
1033  | 
proof cases  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
1034  | 
assume "finite C"  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
1035  | 
from setprod_UN_disjoint[OF this assms]  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
1036  | 
show ?thesis  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
1037  | 
by (simp add: SUP_def)  | 
| 
 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 
hoelzl 
parents: 
44921 
diff
changeset
 | 
1038  | 
qed (force dest: finite_UnionD simp add: setprod_def)  | 
| 15402 | 1039  | 
|
1040  | 
lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==>  | 
|
| 16550 | 1041  | 
(\<Prod>x\<in>A. (\<Prod>y\<in> B x. f x y)) =  | 
| 17189 | 1042  | 
(\<Prod>(x,y)\<in>(SIGMA x:A. B x). f x y)"  | 
| 41550 | 1043  | 
by(simp add:setprod_def fold_image_Sigma split_def)  | 
| 15402 | 1044  | 
|
| 
15409
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
1045  | 
text{*Here we can eliminate the finiteness assumptions, by cases.*}
 | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
1046  | 
lemma setprod_cartesian_product:  | 
| 17189 | 1047  | 
"(\<Prod>x\<in>A. (\<Prod>y\<in> B. f x y)) = (\<Prod>(x,y)\<in>(A <*> B). f x y)"  | 
| 
15409
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
1048  | 
apply (cases "finite A")  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
1049  | 
apply (cases "finite B")  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
1050  | 
apply (simp add: setprod_Sigma)  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
1051  | 
 apply (cases "A={}", simp)
 | 
| 48849 | 1052  | 
apply (simp)  | 
| 
15409
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
1053  | 
apply (auto simp add: setprod_def  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
1054  | 
dest: finite_cartesian_productD1 finite_cartesian_productD2)  | 
| 
 
a063687d24eb
new and stronger lemmas and improved simplification for finite sets
 
paulson 
parents: 
15402 
diff
changeset
 | 
1055  | 
done  | 
| 15402 | 1056  | 
|
| 48861 | 1057  | 
lemma setprod_timesf: "setprod (%x. f x * g x) A = (setprod f A * setprod g A)"  | 
1058  | 
by (fact setprod.F_fun_f)  | 
|
| 15402 | 1059  | 
|
1060  | 
||
1061  | 
subsubsection {* Properties in more restricted classes of structures *}
 | 
|
1062  | 
||
1063  | 
lemma setprod_eq_1_iff [simp]:  | 
|
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1064  | 
"finite F ==> (setprod f F = 1) = (ALL a:F. f a = (1::nat))"  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1065  | 
by (induct set: finite) auto  | 
| 15402 | 1066  | 
|
1067  | 
lemma setprod_zero:  | 
|
| 23277 | 1068  | 
"finite A ==> EX x: A. f x = (0::'a::comm_semiring_1) ==> setprod f A = 0"  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1069  | 
apply (induct set: finite, force, clarsimp)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1070  | 
apply (erule disjE, auto)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1071  | 
done  | 
| 15402 | 1072  | 
|
1073  | 
lemma setprod_nonneg [rule_format]:  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
1074  | 
"(ALL x: A. (0::'a::linordered_semidom) \<le> f x) --> 0 \<le> setprod f A"  | 
| 
30841
 
0813afc97522
generalized setprod_nonneg and setprod_pos to ordered_semidom, simplified proofs
 
huffman 
parents: 
30729 
diff
changeset
 | 
1075  | 
by (cases "finite A", induct set: finite, simp_all add: mult_nonneg_nonneg)  | 
| 
 
0813afc97522
generalized setprod_nonneg and setprod_pos to ordered_semidom, simplified proofs
 
huffman 
parents: 
30729 
diff
changeset
 | 
1076  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
1077  | 
lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::linordered_semidom) < f x)  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1078  | 
--> 0 < setprod f A"  | 
| 
30841
 
0813afc97522
generalized setprod_nonneg and setprod_pos to ordered_semidom, simplified proofs
 
huffman 
parents: 
30729 
diff
changeset
 | 
1079  | 
by (cases "finite A", induct set: finite, simp_all add: mult_pos_pos)  | 
| 15402 | 1080  | 
|
| 30843 | 1081  | 
lemma setprod_zero_iff[simp]: "finite A ==>  | 
1082  | 
  (setprod f A = (0::'a::{comm_semiring_1,no_zero_divisors})) =
 | 
|
1083  | 
(EX x: A. f x = 0)"  | 
|
1084  | 
by (erule finite_induct, auto simp:no_zero_divisors)  | 
|
1085  | 
||
1086  | 
lemma setprod_pos_nat:  | 
|
1087  | 
"finite S ==> (ALL x : S. f x > (0::nat)) ==> setprod f S > 0"  | 
|
1088  | 
using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])  | 
|
| 15402 | 1089  | 
|
| 30863 | 1090  | 
lemma setprod_pos_nat_iff[simp]:  | 
1091  | 
"finite S ==> (setprod f S > 0) = (ALL x : S. f x > (0::nat))"  | 
|
1092  | 
using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])  | 
|
1093  | 
||
| 15402 | 1094  | 
lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \<noteq> 0) ==>  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1095  | 
  (setprod f (A Un B) :: 'a ::{field})
 | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1096  | 
= setprod f A * setprod f B / setprod f (A Int B)"  | 
| 30843 | 1097  | 
by (subst setprod_Un_Int [symmetric], auto)  | 
| 15402 | 1098  | 
|
| 49715 | 1099  | 
lemma setprod_Un2:  | 
1100  | 
assumes "finite (A \<union> B)"  | 
|
1101  | 
shows "setprod f (A \<union> B) = setprod f (A - B) * setprod f (B - A) * setprod f (A \<inter> B)"  | 
|
1102  | 
proof -  | 
|
1103  | 
have "A \<union> B = A - B \<union> (B - A) \<union> A \<inter> B"  | 
|
1104  | 
by auto  | 
|
1105  | 
with assms show ?thesis by simp (subst setprod_Un_disjoint, auto)+  | 
|
1106  | 
qed  | 
|
1107  | 
||
| 15402 | 1108  | 
lemma setprod_diff1: "finite A ==> f a \<noteq> 0 ==>  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1109  | 
  (setprod f (A - {a}) :: 'a :: {field}) =
 | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1110  | 
(if a:A then setprod f A / f a else setprod f A)"  | 
| 36303 | 1111  | 
by (erule finite_induct) (auto simp add: insert_Diff_if)  | 
| 15402 | 1112  | 
|
| 
31906
 
b41d61c768e2
Removed unnecessary conditions concerning nonzero divisors
 
paulson 
parents: 
31465 
diff
changeset
 | 
1113  | 
lemma setprod_inversef:  | 
| 36409 | 1114  | 
fixes f :: "'b \<Rightarrow> 'a::field_inverse_zero"  | 
| 
31906
 
b41d61c768e2
Removed unnecessary conditions concerning nonzero divisors
 
paulson 
parents: 
31465 
diff
changeset
 | 
1115  | 
shows "finite A ==> setprod (inverse \<circ> f) A = inverse (setprod f A)"  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1116  | 
by (erule finite_induct) auto  | 
| 15402 | 1117  | 
|
1118  | 
lemma setprod_dividef:  | 
|
| 36409 | 1119  | 
fixes f :: "'b \<Rightarrow> 'a::field_inverse_zero"  | 
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1120  | 
shows "finite A  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1121  | 
==> setprod (%x. f x / g x) A = setprod f A / setprod g A"  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1122  | 
apply (subgoal_tac  | 
| 15402 | 1123  | 
"setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \<circ> g) x) A")  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1124  | 
apply (erule ssubst)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1125  | 
apply (subst divide_inverse)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1126  | 
apply (subst setprod_timesf)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1127  | 
apply (subst setprod_inversef, assumption+, rule refl)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1128  | 
apply (rule setprod_cong, rule refl)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1129  | 
apply (subst divide_inverse, auto)  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1130  | 
done  | 
| 
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1131  | 
|
| 29925 | 1132  | 
lemma setprod_dvd_setprod [rule_format]:  | 
1133  | 
"(ALL x : A. f x dvd g x) \<longrightarrow> setprod f A dvd setprod g A"  | 
|
1134  | 
apply (cases "finite A")  | 
|
1135  | 
apply (induct set: finite)  | 
|
1136  | 
apply (auto simp add: dvd_def)  | 
|
1137  | 
apply (rule_tac x = "k * ka" in exI)  | 
|
1138  | 
apply (simp add: algebra_simps)  | 
|
1139  | 
done  | 
|
1140  | 
||
1141  | 
lemma setprod_dvd_setprod_subset:  | 
|
1142  | 
"finite B \<Longrightarrow> A <= B \<Longrightarrow> setprod f A dvd setprod f B"  | 
|
1143  | 
apply (subgoal_tac "setprod f B = setprod f A * setprod f (B - A)")  | 
|
1144  | 
apply (unfold dvd_def, blast)  | 
|
1145  | 
apply (subst setprod_Un_disjoint [symmetric])  | 
|
1146  | 
apply (auto elim: finite_subset intro: setprod_cong)  | 
|
1147  | 
done  | 
|
1148  | 
||
1149  | 
lemma setprod_dvd_setprod_subset2:  | 
|
1150  | 
"finite B \<Longrightarrow> A <= B \<Longrightarrow> ALL x : A. (f x::'a::comm_semiring_1) dvd g x \<Longrightarrow>  | 
|
1151  | 
setprod f A dvd setprod g B"  | 
|
1152  | 
apply (rule dvd_trans)  | 
|
1153  | 
apply (rule setprod_dvd_setprod, erule (1) bspec)  | 
|
1154  | 
apply (erule (1) setprod_dvd_setprod_subset)  | 
|
1155  | 
done  | 
|
1156  | 
||
1157  | 
lemma dvd_setprod: "finite A \<Longrightarrow> i:A \<Longrightarrow>  | 
|
1158  | 
(f i ::'a::comm_semiring_1) dvd setprod f A"  | 
|
1159  | 
by (induct set: finite) (auto intro: dvd_mult)  | 
|
1160  | 
||
1161  | 
lemma dvd_setsum [rule_format]: "(ALL i : A. d dvd f i) \<longrightarrow>  | 
|
1162  | 
(d::'a::comm_semiring_1) dvd (SUM x : A. f x)"  | 
|
1163  | 
apply (cases "finite A")  | 
|
1164  | 
apply (induct set: finite)  | 
|
1165  | 
apply auto  | 
|
1166  | 
done  | 
|
1167  | 
||
| 
35171
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1168  | 
lemma setprod_mono:  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1169  | 
fixes f :: "'a \<Rightarrow> 'b\<Colon>linordered_semidom"  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1170  | 
assumes "\<forall>i\<in>A. 0 \<le> f i \<and> f i \<le> g i"  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1171  | 
shows "setprod f A \<le> setprod g A"  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1172  | 
proof (cases "finite A")  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1173  | 
case True  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1174  | 
hence ?thesis "setprod f A \<ge> 0" using subset_refl[of A]  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1175  | 
proof (induct A rule: finite_subset_induct)  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1176  | 
case (insert a F)  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1177  | 
thus "setprod f (insert a F) \<le> setprod g (insert a F)" "0 \<le> setprod f (insert a F)"  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1178  | 
unfolding setprod_insert[OF insert(1,3)]  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1179  | 
using assms[rule_format,OF insert(2)] insert  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1180  | 
by (auto intro: mult_mono mult_nonneg_nonneg)  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1181  | 
qed auto  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1182  | 
thus ?thesis by simp  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1183  | 
qed auto  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1184  | 
|
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1185  | 
lemma abs_setprod:  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1186  | 
  fixes f :: "'a \<Rightarrow> 'b\<Colon>{linordered_field,abs}"
 | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1187  | 
shows "abs (setprod f A) = setprod (\<lambda>x. abs (f x)) A"  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1188  | 
proof (cases "finite A")  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1189  | 
case True thus ?thesis  | 
| 35216 | 1190  | 
by induct (auto simp add: field_simps abs_mult)  | 
| 
35171
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1191  | 
qed auto  | 
| 
 
28f824c7addc
Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
 
hoelzl 
parents: 
35115 
diff
changeset
 | 
1192  | 
|
| 31017 | 1193  | 
lemma setprod_constant: "finite A ==> (\<Prod>x\<in> A. (y::'a::{comm_monoid_mult})) = y^(card A)"
 | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1194  | 
apply (erule finite_induct)  | 
| 35216 | 1195  | 
apply auto  | 
| 
28853
 
69eb69659bf3
Added new fold operator and renamed the old oe to fold_image.
 
nipkow 
parents: 
28823 
diff
changeset
 | 
1196  | 
done  | 
| 15402 | 1197  | 
|
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1198  | 
lemma setprod_gen_delta:  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1199  | 
assumes fS: "finite S"  | 
| 31017 | 1200  | 
  shows "setprod (\<lambda>k. if k=a then b k else c) S = (if a \<in> S then (b a ::'a::{comm_monoid_mult}) * c^ (card S - 1) else c^ card S)"
 | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1201  | 
proof-  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1202  | 
let ?f = "(\<lambda>k. if k=a then b k else c)"  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1203  | 
  {assume a: "a \<notin> S"
 | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1204  | 
hence "\<forall> k\<in> S. ?f k = c" by simp  | 
| 48849 | 1205  | 
hence ?thesis using a setprod_constant[OF fS, of c] by simp }  | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1206  | 
moreover  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1207  | 
  {assume a: "a \<in> S"
 | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1208  | 
    let ?A = "S - {a}"
 | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1209  | 
    let ?B = "{a}"
 | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1210  | 
have eq: "S = ?A \<union> ?B" using a by blast  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1211  | 
    have dj: "?A \<inter> ?B = {}" by simp
 | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1212  | 
from fS have fAB: "finite ?A" "finite ?B" by auto  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1213  | 
have fA0:"setprod ?f ?A = setprod (\<lambda>i. c) ?A"  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1214  | 
apply (rule setprod_cong) by auto  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1215  | 
have cA: "card ?A = card S - 1" using fS a by auto  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1216  | 
have fA1: "setprod ?f ?A = c ^ card ?A" unfolding fA0 apply (rule setprod_constant) using fS by auto  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1217  | 
have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1218  | 
using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1219  | 
by simp  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1220  | 
then have ?thesis using a cA  | 
| 36349 | 1221  | 
by (simp add: fA1 field_simps cong add: setprod_cong cong del: if_weak_cong)}  | 
| 
29674
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1222  | 
ultimately show ?thesis by blast  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1223  | 
qed  | 
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1224  | 
|
| 
 
3857d7eba390
Added theorems setsum_reindex_nonzero, setsum_mono_zero_left, setsum_mono_zero_right, setsum_mono_zero_cong_left, setsum_mono_zero_cong_right, setsum_delta, strong_setprod_reindex_cong, setprod_delta
 
chaieb 
parents: 
29609 
diff
changeset
 | 
1225  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1226  | 
subsection {* Versions of @{const inf} and @{const sup} on non-empty sets *}
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1227  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1228  | 
no_notation times (infixl "*" 70)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1229  | 
no_notation Groups.one ("1")
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1230  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1231  | 
locale semilattice_big = semilattice +  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1232  | 
fixes F :: "'a set \<Rightarrow> 'a"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1233  | 
assumes F_eq: "finite A \<Longrightarrow> F A = fold1 (op *) A"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1234  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1235  | 
sublocale semilattice_big < folding_one_idem proof  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1236  | 
qed (simp_all add: F_eq)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1237  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1238  | 
notation times (infixl "*" 70)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1239  | 
notation Groups.one ("1")
 | 
| 22917 | 1240  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1241  | 
context lattice  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1242  | 
begin  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1243  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1244  | 
definition Inf_fin :: "'a set \<Rightarrow> 'a" ("\<Sqinter>\<^bsub>fin\<^esub>_" [900] 900) where
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1245  | 
"Inf_fin = fold1 inf"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1246  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1247  | 
definition Sup_fin :: "'a set \<Rightarrow> 'a" ("\<Squnion>\<^bsub>fin\<^esub>_" [900] 900) where
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1248  | 
"Sup_fin = fold1 sup"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1249  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1250  | 
end  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1251  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1252  | 
sublocale lattice < Inf_fin!: semilattice_big inf Inf_fin proof  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1253  | 
qed (simp add: Inf_fin_def)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1254  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1255  | 
sublocale lattice < Sup_fin!: semilattice_big sup Sup_fin proof  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1256  | 
qed (simp add: Sup_fin_def)  | 
| 22917 | 1257  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
1258  | 
context semilattice_inf  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1259  | 
begin  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1260  | 
|
| 
36635
 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 
haftmann 
parents: 
36622 
diff
changeset
 | 
1261  | 
lemma ab_semigroup_idem_mult_inf:  | 
| 
 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 
haftmann 
parents: 
36622 
diff
changeset
 | 
1262  | 
"class.ab_semigroup_idem_mult inf"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1263  | 
proof qed (rule inf_assoc inf_commute inf_idem)+  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1264  | 
|
| 46033 | 1265  | 
lemma fold_inf_insert[simp]: "finite A \<Longrightarrow> Finite_Set.fold inf b (insert a A) = inf a (Finite_Set.fold inf b A)"  | 
| 
42871
 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 
haftmann 
parents: 
42284 
diff
changeset
 | 
1266  | 
by(rule comp_fun_idem.fold_insert_idem[OF ab_semigroup_idem_mult.comp_fun_idem[OF ab_semigroup_idem_mult_inf]])  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1267  | 
|
| 46033 | 1268  | 
lemma inf_le_fold_inf: "finite A \<Longrightarrow> ALL a:A. b \<le> a \<Longrightarrow> inf b c \<le> Finite_Set.fold inf c A"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1269  | 
by (induct pred: finite) (auto intro: le_infI1)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1270  | 
|
| 46033 | 1271  | 
lemma fold_inf_le_inf: "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> Finite_Set.fold inf b A \<le> inf a b"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1272  | 
proof(induct arbitrary: a pred:finite)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1273  | 
case empty thus ?case by simp  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1274  | 
next  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1275  | 
case (insert x A)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1276  | 
show ?case  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1277  | 
proof cases  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1278  | 
    assume "A = {}" thus ?thesis using insert by simp
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1279  | 
next  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1280  | 
    assume "A \<noteq> {}" thus ?thesis using insert by (auto intro: le_infI2)
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1281  | 
qed  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1282  | 
qed  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1283  | 
|
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1284  | 
lemma below_fold1_iff:  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1285  | 
  assumes "finite A" "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1286  | 
shows "x \<le> fold1 inf A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1287  | 
proof -  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1288  | 
interpret ab_semigroup_idem_mult inf  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1289  | 
by (rule ab_semigroup_idem_mult_inf)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1290  | 
show ?thesis using assms by (induct rule: finite_ne_induct) simp_all  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1291  | 
qed  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1292  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1293  | 
lemma fold1_belowI:  | 
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1294  | 
assumes "finite A"  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1295  | 
and "a \<in> A"  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1296  | 
shows "fold1 inf A \<le> a"  | 
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1297  | 
proof -  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1298  | 
  from assms have "A \<noteq> {}" by auto
 | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1299  | 
  from `finite A` `A \<noteq> {}` `a \<in> A` show ?thesis
 | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1300  | 
proof (induct rule: finite_ne_induct)  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1301  | 
case singleton thus ?case by simp  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1302  | 
next  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1303  | 
interpret ab_semigroup_idem_mult inf  | 
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1304  | 
by (rule ab_semigroup_idem_mult_inf)  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1305  | 
case (insert x F)  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1306  | 
from insert(5) have "a = x \<or> a \<in> F" by simp  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1307  | 
thus ?case  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1308  | 
proof  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1309  | 
assume "a = x" thus ?thesis using insert  | 
| 29667 | 1310  | 
by (simp add: mult_ac)  | 
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1311  | 
next  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1312  | 
assume "a \<in> F"  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1313  | 
hence bel: "fold1 inf F \<le> a" by (rule insert)  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1314  | 
have "inf (fold1 inf (insert x F)) a = inf x (inf (fold1 inf F) a)"  | 
| 29667 | 1315  | 
using insert by (simp add: mult_ac)  | 
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1316  | 
also have "inf (fold1 inf F) a = fold1 inf F"  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1317  | 
using bel by (auto intro: antisym)  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1318  | 
also have "inf x \<dots> = fold1 inf (insert x F)"  | 
| 29667 | 1319  | 
using insert by (simp add: mult_ac)  | 
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1320  | 
finally have aux: "inf (fold1 inf (insert x F)) a = fold1 inf (insert x F)" .  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1321  | 
moreover have "inf (fold1 inf (insert x F)) a \<le> a" by simp  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1322  | 
ultimately show ?thesis by simp  | 
| 
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1323  | 
qed  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1324  | 
qed  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1325  | 
qed  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1326  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1327  | 
end  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1328  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1329  | 
context semilattice_sup  | 
| 22917 | 1330  | 
begin  | 
1331  | 
||
| 
36635
 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 
haftmann 
parents: 
36622 
diff
changeset
 | 
1332  | 
lemma ab_semigroup_idem_mult_sup: "class.ab_semigroup_idem_mult sup"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1333  | 
by (rule semilattice_inf.ab_semigroup_idem_mult_inf)(rule dual_semilattice)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1334  | 
|
| 46033 | 1335  | 
lemma fold_sup_insert[simp]: "finite A \<Longrightarrow> Finite_Set.fold sup b (insert a A) = sup a (Finite_Set.fold sup b A)"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1336  | 
by(rule semilattice_inf.fold_inf_insert)(rule dual_semilattice)  | 
| 22917 | 1337  | 
|
| 46033 | 1338  | 
lemma fold_sup_le_sup: "finite A \<Longrightarrow> ALL a:A. a \<le> b \<Longrightarrow> Finite_Set.fold sup c A \<le> sup b c"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1339  | 
by(rule semilattice_inf.inf_le_fold_inf)(rule dual_semilattice)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1340  | 
|
| 46033 | 1341  | 
lemma sup_le_fold_sup: "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> sup a b \<le> Finite_Set.fold sup b A"  | 
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1342  | 
by(rule semilattice_inf.fold_inf_le_inf)(rule dual_semilattice)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1343  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1344  | 
end  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1345  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1346  | 
context lattice  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1347  | 
begin  | 
| 25062 | 1348  | 
|
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1349  | 
lemma Inf_le_Sup [simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> \<Sqinter>\<^bsub>fin\<^esub>A \<le> \<Squnion>\<^bsub>fin\<^esub>A"
 | 
| 24342 | 1350  | 
apply(unfold Sup_fin_def Inf_fin_def)  | 
| 15500 | 1351  | 
apply(subgoal_tac "EX a. a:A")  | 
1352  | 
prefer 2 apply blast  | 
|
1353  | 
apply(erule exE)  | 
|
| 22388 | 1354  | 
apply(rule order_trans)  | 
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1355  | 
apply(erule (1) fold1_belowI)  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
1356  | 
apply(erule (1) semilattice_inf.fold1_belowI [OF dual_semilattice])  | 
| 15500 | 1357  | 
done  | 
1358  | 
||
| 24342 | 1359  | 
lemma sup_Inf_absorb [simp]:  | 
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1360  | 
"finite A \<Longrightarrow> a \<in> A \<Longrightarrow> sup a (\<Sqinter>\<^bsub>fin\<^esub>A) = a"  | 
| 
15512
 
ed1fa4617f52
Extracted generic lattice stuff to new Lattice_Locales.thy
 
nipkow 
parents: 
15510 
diff
changeset
 | 
1361  | 
apply(subst sup_commute)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1362  | 
apply(simp add: Inf_fin_def sup_absorb2 fold1_belowI)  | 
| 15504 | 1363  | 
done  | 
1364  | 
||
| 24342 | 1365  | 
lemma inf_Sup_absorb [simp]:  | 
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1366  | 
"finite A \<Longrightarrow> a \<in> A \<Longrightarrow> inf a (\<Squnion>\<^bsub>fin\<^esub>A) = a"  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1367  | 
by (simp add: Sup_fin_def inf_absorb1  | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
1368  | 
semilattice_inf.fold1_belowI [OF dual_semilattice])  | 
| 24342 | 1369  | 
|
1370  | 
end  | 
|
1371  | 
||
1372  | 
context distrib_lattice  | 
|
1373  | 
begin  | 
|
1374  | 
||
1375  | 
lemma sup_Inf1_distrib:  | 
|
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1376  | 
assumes "finite A"  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1377  | 
    and "A \<noteq> {}"
 | 
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1378  | 
  shows "sup x (\<Sqinter>\<^bsub>fin\<^esub>A) = \<Sqinter>\<^bsub>fin\<^esub>{sup x a|a. a \<in> A}"
 | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1379  | 
proof -  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1380  | 
interpret ab_semigroup_idem_mult inf  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1381  | 
by (rule ab_semigroup_idem_mult_inf)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1382  | 
from assms show ?thesis  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1383  | 
by (simp add: Inf_fin_def image_def  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1384  | 
hom_fold1_commute [where h="sup x", OF sup_inf_distrib1])  | 
| 26792 | 1385  | 
(rule arg_cong [where f="fold1 inf"], blast)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1386  | 
qed  | 
| 18423 | 1387  | 
|
| 24342 | 1388  | 
lemma sup_Inf2_distrib:  | 
1389  | 
  assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
 | 
|
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1390  | 
  shows "sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B) = \<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B}"
 | 
| 24342 | 1391  | 
using A proof (induct rule: finite_ne_induct)  | 
| 15500 | 1392  | 
case singleton thus ?case  | 
| 41550 | 1393  | 
by (simp add: sup_Inf1_distrib [OF B])  | 
| 15500 | 1394  | 
next  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1395  | 
interpret ab_semigroup_idem_mult inf  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1396  | 
by (rule ab_semigroup_idem_mult_inf)  | 
| 15500 | 1397  | 
case (insert x A)  | 
| 25062 | 1398  | 
  have finB: "finite {sup x b |b. b \<in> B}"
 | 
1399  | 
by(rule finite_surj[where f = "sup x", OF B(1)], auto)  | 
|
1400  | 
  have finAB: "finite {sup a b |a b. a \<in> A \<and> b \<in> B}"
 | 
|
| 15500 | 1401  | 
proof -  | 
| 25062 | 1402  | 
    have "{sup a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {sup a b})"
 | 
| 15500 | 1403  | 
by blast  | 
| 15517 | 1404  | 
thus ?thesis by(simp add: insert(1) B(1))  | 
| 15500 | 1405  | 
qed  | 
| 25062 | 1406  | 
  have ne: "{sup a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
 | 
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1407  | 
have "sup (\<Sqinter>\<^bsub>fin\<^esub>(insert x A)) (\<Sqinter>\<^bsub>fin\<^esub>B) = sup (inf x (\<Sqinter>\<^bsub>fin\<^esub>A)) (\<Sqinter>\<^bsub>fin\<^esub>B)"  | 
| 41550 | 1408  | 
using insert by simp  | 
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1409  | 
also have "\<dots> = inf (sup x (\<Sqinter>\<^bsub>fin\<^esub>B)) (sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B))" by(rule sup_inf_distrib2)  | 
| 
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1410  | 
  also have "\<dots> = inf (\<Sqinter>\<^bsub>fin\<^esub>{sup x b|b. b \<in> B}) (\<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B})"
 | 
| 15500 | 1411  | 
using insert by(simp add:sup_Inf1_distrib[OF B])  | 
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1412  | 
  also have "\<dots> = \<Sqinter>\<^bsub>fin\<^esub>({sup x b |b. b \<in> B} \<union> {sup a b |a b. a \<in> A \<and> b \<in> B})"
 | 
| 
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1413  | 
(is "_ = \<Sqinter>\<^bsub>fin\<^esub>?M")  | 
| 15500 | 1414  | 
using B insert  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1415  | 
by (simp add: Inf_fin_def fold1_Un2 [OF finB _ finAB ne])  | 
| 25062 | 1416  | 
  also have "?M = {sup a b |a b. a \<in> insert x A \<and> b \<in> B}"
 | 
| 15500 | 1417  | 
by blast  | 
1418  | 
finally show ?case .  | 
|
1419  | 
qed  | 
|
1420  | 
||
| 24342 | 1421  | 
lemma inf_Sup1_distrib:  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1422  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1423  | 
  shows "inf x (\<Squnion>\<^bsub>fin\<^esub>A) = \<Squnion>\<^bsub>fin\<^esub>{inf x a|a. a \<in> A}"
 | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1424  | 
proof -  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1425  | 
interpret ab_semigroup_idem_mult sup  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1426  | 
by (rule ab_semigroup_idem_mult_sup)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1427  | 
from assms show ?thesis  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1428  | 
by (simp add: Sup_fin_def image_def hom_fold1_commute [where h="inf x", OF inf_sup_distrib1])  | 
| 26792 | 1429  | 
(rule arg_cong [where f="fold1 sup"], blast)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1430  | 
qed  | 
| 18423 | 1431  | 
|
| 24342 | 1432  | 
lemma inf_Sup2_distrib:  | 
1433  | 
  assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
 | 
|
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1434  | 
  shows "inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B) = \<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B}"
 | 
| 24342 | 1435  | 
using A proof (induct rule: finite_ne_induct)  | 
| 18423 | 1436  | 
case singleton thus ?case  | 
| 44921 | 1437  | 
by(simp add: inf_Sup1_distrib [OF B])  | 
| 18423 | 1438  | 
next  | 
1439  | 
case (insert x A)  | 
|
| 25062 | 1440  | 
  have finB: "finite {inf x b |b. b \<in> B}"
 | 
1441  | 
by(rule finite_surj[where f = "%b. inf x b", OF B(1)], auto)  | 
|
1442  | 
  have finAB: "finite {inf a b |a b. a \<in> A \<and> b \<in> B}"
 | 
|
| 18423 | 1443  | 
proof -  | 
| 25062 | 1444  | 
    have "{inf a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {inf a b})"
 | 
| 18423 | 1445  | 
by blast  | 
1446  | 
thus ?thesis by(simp add: insert(1) B(1))  | 
|
1447  | 
qed  | 
|
| 25062 | 1448  | 
  have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
 | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1449  | 
interpret ab_semigroup_idem_mult sup  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1450  | 
by (rule ab_semigroup_idem_mult_sup)  | 
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1451  | 
have "inf (\<Squnion>\<^bsub>fin\<^esub>(insert x A)) (\<Squnion>\<^bsub>fin\<^esub>B) = inf (sup x (\<Squnion>\<^bsub>fin\<^esub>A)) (\<Squnion>\<^bsub>fin\<^esub>B)"  | 
| 41550 | 1452  | 
using insert by simp  | 
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1453  | 
also have "\<dots> = sup (inf x (\<Squnion>\<^bsub>fin\<^esub>B)) (inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B))" by(rule inf_sup_distrib2)  | 
| 
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1454  | 
  also have "\<dots> = sup (\<Squnion>\<^bsub>fin\<^esub>{inf x b|b. b \<in> B}) (\<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B})"
 | 
| 18423 | 1455  | 
using insert by(simp add:inf_Sup1_distrib[OF B])  | 
| 
31916
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1456  | 
  also have "\<dots> = \<Squnion>\<^bsub>fin\<^esub>({inf x b |b. b \<in> B} \<union> {inf a b |a b. a \<in> A \<and> b \<in> B})"
 | 
| 
 
f3227bb306a4
recovered subscripts, which were lost in b41d61c768e2 (due to Emacs accident?);
 
wenzelm 
parents: 
31907 
diff
changeset
 | 
1457  | 
(is "_ = \<Squnion>\<^bsub>fin\<^esub>?M")  | 
| 18423 | 1458  | 
using B insert  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1459  | 
by (simp add: Sup_fin_def fold1_Un2 [OF finB _ finAB ne])  | 
| 25062 | 1460  | 
  also have "?M = {inf a b |a b. a \<in> insert x A \<and> b \<in> B}"
 | 
| 18423 | 1461  | 
by blast  | 
1462  | 
finally show ?case .  | 
|
1463  | 
qed  | 
|
1464  | 
||
| 24342 | 1465  | 
end  | 
1466  | 
||
| 
35719
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1467  | 
context complete_lattice  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1468  | 
begin  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1469  | 
|
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1470  | 
lemma Inf_fin_Inf:  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1471  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1472  | 
shows "\<Sqinter>\<^bsub>fin\<^esub>A = Inf A"  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1473  | 
proof -  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1474  | 
interpret ab_semigroup_idem_mult inf  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1475  | 
by (rule ab_semigroup_idem_mult_inf)  | 
| 44918 | 1476  | 
  from `A \<noteq> {}` obtain b B where "A = {b} \<union> B" by auto
 | 
| 
35719
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1477  | 
moreover with `finite A` have "finite B" by simp  | 
| 44918 | 1478  | 
ultimately show ?thesis  | 
1479  | 
by (simp add: Inf_fin_def fold1_eq_fold_idem inf_Inf_fold_inf [symmetric])  | 
|
| 
35719
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1480  | 
qed  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1481  | 
|
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1482  | 
lemma Sup_fin_Sup:  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1483  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1484  | 
shows "\<Squnion>\<^bsub>fin\<^esub>A = Sup A"  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1485  | 
proof -  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1486  | 
interpret ab_semigroup_idem_mult sup  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1487  | 
by (rule ab_semigroup_idem_mult_sup)  | 
| 44918 | 1488  | 
  from `A \<noteq> {}` obtain b B where "A = {b} \<union> B" by auto
 | 
| 
35719
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1489  | 
moreover with `finite A` have "finite B" by simp  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1490  | 
ultimately show ?thesis  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1491  | 
by (simp add: Sup_fin_def fold1_eq_fold_idem sup_Sup_fold_sup [symmetric])  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1492  | 
qed  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1493  | 
|
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1494  | 
end  | 
| 
 
99b6152aedf5
split off theory Big_Operators from theory Finite_Set
 
haftmann 
parents: 
35577 
diff
changeset
 | 
1495  | 
|
| 22917 | 1496  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1497  | 
subsection {* Versions of @{const min} and @{const max} on non-empty sets *}
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1498  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1499  | 
definition (in linorder) Min :: "'a set \<Rightarrow> 'a" where  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1500  | 
"Min = fold1 min"  | 
| 22917 | 1501  | 
|
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1502  | 
definition (in linorder) Max :: "'a set \<Rightarrow> 'a" where  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1503  | 
"Max = fold1 max"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1504  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1505  | 
sublocale linorder < Min!: semilattice_big min Min proof  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1506  | 
qed (simp add: Min_def)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1507  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1508  | 
sublocale linorder < Max!: semilattice_big max Max proof  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1509  | 
qed (simp add: Max_def)  | 
| 22917 | 1510  | 
|
| 24342 | 1511  | 
context linorder  | 
| 22917 | 1512  | 
begin  | 
1513  | 
||
| 
35816
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1514  | 
lemmas Min_singleton = Min.singleton  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1515  | 
lemmas Max_singleton = Max.singleton  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1516  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1517  | 
lemma Min_insert:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1518  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1519  | 
shows "Min (insert x A) = min x (Min A)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1520  | 
using assms by simp  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1521  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1522  | 
lemma Max_insert:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1523  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1524  | 
shows "Max (insert x A) = max x (Max A)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1525  | 
using assms by simp  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1526  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1527  | 
lemma Min_Un:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1528  | 
  assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1529  | 
shows "Min (A \<union> B) = min (Min A) (Min B)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1530  | 
using assms by (rule Min.union_idem)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1531  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1532  | 
lemma Max_Un:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1533  | 
  assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1534  | 
shows "Max (A \<union> B) = max (Max A) (Max B)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1535  | 
using assms by (rule Max.union_idem)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1536  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1537  | 
lemma hom_Min_commute:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1538  | 
assumes "\<And>x y. h (min x y) = min (h x) (h y)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1539  | 
    and "finite N" and "N \<noteq> {}"
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1540  | 
shows "h (Min N) = Min (h ` N)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1541  | 
using assms by (rule Min.hom_commute)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1542  | 
|
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1543  | 
lemma hom_Max_commute:  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1544  | 
assumes "\<And>x y. h (max x y) = max (h x) (h y)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1545  | 
    and "finite N" and "N \<noteq> {}"
 | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1546  | 
shows "h (Max N) = Max (h ` N)"  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1547  | 
using assms by (rule Max.hom_commute)  | 
| 
 
2449e026483d
generic locale for big operators in monoids; dropped odd interpretation of comm_monoid_mult into comm_monoid_add
 
haftmann 
parents: 
35722 
diff
changeset
 | 
1548  | 
|
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1549  | 
lemma ab_semigroup_idem_mult_min:  | 
| 
36635
 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 
haftmann 
parents: 
36622 
diff
changeset
 | 
1550  | 
"class.ab_semigroup_idem_mult min"  | 
| 28823 | 1551  | 
proof qed (auto simp add: min_def)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1552  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1553  | 
lemma ab_semigroup_idem_mult_max:  | 
| 
36635
 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 
haftmann 
parents: 
36622 
diff
changeset
 | 
1554  | 
"class.ab_semigroup_idem_mult max"  | 
| 28823 | 1555  | 
proof qed (auto simp add: max_def)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1556  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1557  | 
lemma max_lattice:  | 
| 44845 | 1558  | 
"class.semilattice_inf max (op \<ge>) (op >)"  | 
| 
32203
 
992ac8942691
adapted to localized interpretation of min/max-lattice
 
haftmann 
parents: 
32075 
diff
changeset
 | 
1559  | 
by (fact min_max.dual_semilattice)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1560  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1561  | 
lemma dual_max:  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1562  | 
"ord.max (op \<ge>) = min"  | 
| 46904 | 1563  | 
by (auto simp add: ord.max_def min_def fun_eq_iff)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1564  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1565  | 
lemma dual_min:  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1566  | 
"ord.min (op \<ge>) = max"  | 
| 46904 | 1567  | 
by (auto simp add: ord.min_def max_def fun_eq_iff)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1568  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1569  | 
lemma strict_below_fold1_iff:  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1570  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1571  | 
shows "x < fold1 min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1572  | 
proof -  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1573  | 
interpret ab_semigroup_idem_mult min  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1574  | 
by (rule ab_semigroup_idem_mult_min)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1575  | 
from assms show ?thesis  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1576  | 
by (induct rule: finite_ne_induct)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1577  | 
(simp_all add: fold1_insert)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1578  | 
qed  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1579  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1580  | 
lemma fold1_below_iff:  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1581  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1582  | 
shows "fold1 min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1583  | 
proof -  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1584  | 
interpret ab_semigroup_idem_mult min  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1585  | 
by (rule ab_semigroup_idem_mult_min)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1586  | 
from assms show ?thesis  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1587  | 
by (induct rule: finite_ne_induct)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1588  | 
(simp_all add: fold1_insert min_le_iff_disj)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1589  | 
qed  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1590  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1591  | 
lemma fold1_strict_below_iff:  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1592  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1593  | 
shows "fold1 min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1594  | 
proof -  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1595  | 
interpret ab_semigroup_idem_mult min  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1596  | 
by (rule ab_semigroup_idem_mult_min)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1597  | 
from assms show ?thesis  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1598  | 
by (induct rule: finite_ne_induct)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1599  | 
(simp_all add: fold1_insert min_less_iff_disj)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1600  | 
qed  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1601  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1602  | 
lemma fold1_antimono:  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1603  | 
  assumes "A \<noteq> {}" and "A \<subseteq> B" and "finite B"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1604  | 
shows "fold1 min B \<le> fold1 min A"  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1605  | 
proof cases  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1606  | 
assume "A = B" thus ?thesis by simp  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1607  | 
next  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1608  | 
interpret ab_semigroup_idem_mult min  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1609  | 
by (rule ab_semigroup_idem_mult_min)  | 
| 41550 | 1610  | 
assume neq: "A \<noteq> B"  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1611  | 
have B: "B = A \<union> (B-A)" using `A \<subseteq> B` by blast  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1612  | 
have "fold1 min B = fold1 min (A \<union> (B-A))" by(subst B)(rule refl)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1613  | 
also have "\<dots> = min (fold1 min A) (fold1 min (B-A))"  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1614  | 
proof -  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1615  | 
have "finite A" by(rule finite_subset[OF `A \<subseteq> B` `finite B`])  | 
| 41550 | 1616  | 
moreover have "finite(B-A)" by(rule finite_Diff[OF `finite B`])  | 
1617  | 
    moreover have "(B-A) \<noteq> {}" using assms neq by blast
 | 
|
1618  | 
    moreover have "A Int (B-A) = {}" using assms by blast
 | 
|
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1619  | 
    ultimately show ?thesis using `A \<noteq> {}` by (rule_tac fold1_Un)
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1620  | 
qed  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1621  | 
also have "\<dots> \<le> fold1 min A" by (simp add: min_le_iff_disj)  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1622  | 
finally show ?thesis .  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1623  | 
qed  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1624  | 
|
| 24427 | 1625  | 
lemma Min_in [simp]:  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1626  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1627  | 
shows "Min A \<in> A"  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1628  | 
proof -  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1629  | 
interpret ab_semigroup_idem_mult min  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1630  | 
by (rule ab_semigroup_idem_mult_min)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44845 
diff
changeset
 | 
1631  | 
from assms fold1_in show ?thesis by (fastforce simp: Min_def min_def)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1632  | 
qed  | 
| 15392 | 1633  | 
|
| 24427 | 1634  | 
lemma Max_in [simp]:  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1635  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1636  | 
shows "Max A \<in> A"  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1637  | 
proof -  | 
| 
29509
 
1ff0f3f08a7b
migrated class package to new locale implementation
 
haftmann 
parents: 
29223 
diff
changeset
 | 
1638  | 
interpret ab_semigroup_idem_mult max  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1639  | 
by (rule ab_semigroup_idem_mult_max)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44845 
diff
changeset
 | 
1640  | 
from assms fold1_in [of A] show ?thesis by (fastforce simp: Max_def max_def)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1641  | 
qed  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1642  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1643  | 
lemma Min_le [simp]:  | 
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1644  | 
assumes "finite A" and "x \<in> A"  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1645  | 
shows "Min A \<le> x"  | 
| 
32203
 
992ac8942691
adapted to localized interpretation of min/max-lattice
 
haftmann 
parents: 
32075 
diff
changeset
 | 
1646  | 
using assms by (simp add: Min_def min_max.fold1_belowI)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1647  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1648  | 
lemma Max_ge [simp]:  | 
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1649  | 
assumes "finite A" and "x \<in> A"  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1650  | 
shows "x \<le> Max A"  | 
| 44921 | 1651  | 
by (simp add: Max_def semilattice_inf.fold1_belowI [OF max_lattice] assms)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1652  | 
|
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35722 
diff
changeset
 | 
1653  | 
lemma Min_ge_iff [simp, no_atp]:  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1654  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1655  | 
shows "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"  | 
| 
32203
 
992ac8942691
adapted to localized interpretation of min/max-lattice
 
haftmann 
parents: 
32075 
diff
changeset
 | 
1656  | 
using assms by (simp add: Min_def min_max.below_fold1_iff)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1657  | 
|
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35722 
diff
changeset
 | 
1658  | 
lemma Max_le_iff [simp, no_atp]:  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1659  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1660  | 
shows "Max A \<le> x \<longleftrightarrow> (\<forall>a\<in>A. a \<le> x)"  | 
| 44921 | 1661  | 
by (simp add: Max_def semilattice_inf.below_fold1_iff [OF max_lattice] assms)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1662  | 
|
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35722 
diff
changeset
 | 
1663  | 
lemma Min_gr_iff [simp, no_atp]:  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1664  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1665  | 
shows "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"  | 
| 
32203
 
992ac8942691
adapted to localized interpretation of min/max-lattice
 
haftmann 
parents: 
32075 
diff
changeset
 | 
1666  | 
using assms by (simp add: Min_def strict_below_fold1_iff)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1667  | 
|
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35722 
diff
changeset
 | 
1668  | 
lemma Max_less_iff [simp, no_atp]:  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1669  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1670  | 
shows "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)"  | 
| 44921 | 1671  | 
by (simp add: Max_def linorder.dual_max [OF dual_linorder]  | 
1672  | 
linorder.strict_below_fold1_iff [OF dual_linorder] assms)  | 
|
| 18493 | 1673  | 
|
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35722 
diff
changeset
 | 
1674  | 
lemma Min_le_iff [no_atp]:  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1675  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1676  | 
shows "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"  | 
| 
32203
 
992ac8942691
adapted to localized interpretation of min/max-lattice
 
haftmann 
parents: 
32075 
diff
changeset
 | 
1677  | 
using assms by (simp add: Min_def fold1_below_iff)  | 
| 
15497
 
53bca254719a
Added semi-lattice locales and reorganized fold1 lemmas
 
nipkow 
parents: 
15487 
diff
changeset
 | 
1678  | 
|
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35722 
diff
changeset
 | 
1679  | 
lemma Max_ge_iff [no_atp]:  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1680  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1681  | 
shows "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)"  | 
| 44921 | 1682  | 
by (simp add: Max_def linorder.dual_max [OF dual_linorder]  | 
1683  | 
linorder.fold1_below_iff [OF dual_linorder] assms)  | 
|
| 22917 | 1684  | 
|
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35722 
diff
changeset
 | 
1685  | 
lemma Min_less_iff [no_atp]:  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1686  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1687  | 
shows "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"  | 
| 
32203
 
992ac8942691
adapted to localized interpretation of min/max-lattice
 
haftmann 
parents: 
32075 
diff
changeset
 | 
1688  | 
using assms by (simp add: Min_def fold1_strict_below_iff)  | 
| 22917 | 1689  | 
|
| 
35828
 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 
blanchet 
parents: 
35722 
diff
changeset
 | 
1690  | 
lemma Max_gr_iff [no_atp]:  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1691  | 
  assumes "finite A" and "A \<noteq> {}"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1692  | 
shows "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)"  | 
| 44921 | 1693  | 
by (simp add: Max_def linorder.dual_max [OF dual_linorder]  | 
1694  | 
linorder.fold1_strict_below_iff [OF dual_linorder] assms)  | 
|
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1695  | 
|
| 30325 | 1696  | 
lemma Min_eqI:  | 
1697  | 
assumes "finite A"  | 
|
1698  | 
assumes "\<And>y. y \<in> A \<Longrightarrow> y \<ge> x"  | 
|
1699  | 
and "x \<in> A"  | 
|
1700  | 
shows "Min A = x"  | 
|
1701  | 
proof (rule antisym)  | 
|
1702  | 
  from `x \<in> A` have "A \<noteq> {}" by auto
 | 
|
1703  | 
with assms show "Min A \<ge> x" by simp  | 
|
1704  | 
next  | 
|
1705  | 
from assms show "x \<ge> Min A" by simp  | 
|
1706  | 
qed  | 
|
1707  | 
||
1708  | 
lemma Max_eqI:  | 
|
1709  | 
assumes "finite A"  | 
|
1710  | 
assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"  | 
|
1711  | 
and "x \<in> A"  | 
|
1712  | 
shows "Max A = x"  | 
|
1713  | 
proof (rule antisym)  | 
|
1714  | 
  from `x \<in> A` have "A \<noteq> {}" by auto
 | 
|
1715  | 
with assms show "Max A \<le> x" by simp  | 
|
1716  | 
next  | 
|
1717  | 
from assms show "x \<le> Max A" by simp  | 
|
1718  | 
qed  | 
|
1719  | 
||
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1720  | 
lemma Min_antimono:  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1721  | 
  assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1722  | 
shows "Min N \<le> Min M"  | 
| 
32203
 
992ac8942691
adapted to localized interpretation of min/max-lattice
 
haftmann 
parents: 
32075 
diff
changeset
 | 
1723  | 
using assms by (simp add: Min_def fold1_antimono)  | 
| 
26041
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1724  | 
|
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1725  | 
lemma Max_mono:  | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1726  | 
  assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
 | 
| 
 
c2e15e65165f
locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
 
haftmann 
parents: 
25571 
diff
changeset
 | 
1727  | 
shows "Max M \<le> Max N"  | 
| 44921 | 1728  | 
by (simp add: Max_def linorder.dual_max [OF dual_linorder]  | 
1729  | 
linorder.fold1_antimono [OF dual_linorder] assms)  | 
|
| 22917 | 1730  | 
|
| 32006 | 1731  | 
lemma finite_linorder_max_induct[consumes 1, case_names empty insert]:  | 
| 
36079
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1732  | 
assumes fin: "finite A"  | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1733  | 
 and   empty: "P {}" 
 | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1734  | 
and insert: "(!!b A. finite A \<Longrightarrow> ALL a:A. a < b \<Longrightarrow> P A \<Longrightarrow> P(insert b A))"  | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1735  | 
shows "P A"  | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1736  | 
using fin empty insert  | 
| 32006 | 1737  | 
proof (induct rule: finite_psubset_induct)  | 
| 
36079
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1738  | 
case (psubset A)  | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1739  | 
  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 
 | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1740  | 
have fin: "finite A" by fact  | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1741  | 
  have empty: "P {}" by fact
 | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1742  | 
have step: "\<And>b A. \<lbrakk>finite A; \<forall>a\<in>A. a < b; P A\<rbrakk> \<Longrightarrow> P (insert b A)" by fact  | 
| 
26748
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
26465 
diff
changeset
 | 
1743  | 
show "P A"  | 
| 
26757
 
e775accff967
thms Max_ge, Min_le: dropped superfluous premise
 
haftmann 
parents: 
26748 
diff
changeset
 | 
1744  | 
  proof (cases "A = {}")
 | 
| 
36079
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1745  | 
    assume "A = {}" 
 | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1746  | 
    then show "P A" using `P {}` by simp
 | 
| 
26748
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
26465 
diff
changeset
 | 
1747  | 
next  | 
| 
36079
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1748  | 
    let ?B = "A - {Max A}" 
 | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1749  | 
let ?A = "insert (Max A) ?B"  | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1750  | 
have "finite ?B" using `finite A` by simp  | 
| 
26748
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
26465 
diff
changeset
 | 
1751  | 
    assume "A \<noteq> {}"
 | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
26465 
diff
changeset
 | 
1752  | 
with `finite A` have "Max A : A" by auto  | 
| 
36079
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1753  | 
then have A: "?A = A" using insert_Diff_single insert_absorb by auto  | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1754  | 
    then have "P ?B" using `P {}` step IH[of ?B] by blast
 | 
| 
 
fa0e354e6a39
simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35938 
diff
changeset
 | 
1755  | 
moreover  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44845 
diff
changeset
 | 
1756  | 
have "\<forall>a\<in>?B. a < Max A" using Max_ge [OF `finite A`] by fastforce  | 
| 
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44845 
diff
changeset
 | 
1757  | 
ultimately show "P A" using A insert_Diff_single step[OF `finite ?B`] by fastforce  | 
| 
26748
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
26465 
diff
changeset
 | 
1758  | 
qed  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
26465 
diff
changeset
 | 
1759  | 
qed  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
26465 
diff
changeset
 | 
1760  | 
|
| 32006 | 1761  | 
lemma finite_linorder_min_induct[consumes 1, case_names empty insert]:  | 
| 33434 | 1762  | 
 "\<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"
 | 
| 32006 | 1763  | 
by(rule linorder.finite_linorder_max_induct[OF dual_linorder])  | 
1764  | 
||
| 22917 | 1765  | 
end  | 
1766  | 
||
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34223 
diff
changeset
 | 
1767  | 
context linordered_ab_semigroup_add  | 
| 22917 | 1768  | 
begin  | 
1769  | 
||
1770  | 
lemma add_Min_commute:  | 
|
1771  | 
fixes k  | 
|
| 25062 | 1772  | 
  assumes "finite N" and "N \<noteq> {}"
 | 
1773  | 
  shows "k + Min N = Min {k + m | m. m \<in> N}"
 | 
|
1774  | 
proof -  | 
|
1775  | 
have "\<And>x y. k + min x y = min (k + x) (k + y)"  | 
|
1776  | 
by (simp add: min_def not_le)  | 
|
1777  | 
(blast intro: antisym less_imp_le add_left_mono)  | 
|
1778  | 
with assms show ?thesis  | 
|
1779  | 
using hom_Min_commute [of "plus k" N]  | 
|
1780  | 
by simp (blast intro: arg_cong [where f = Min])  | 
|
1781  | 
qed  | 
|
| 22917 | 1782  | 
|
1783  | 
lemma add_Max_commute:  | 
|
1784  | 
fixes k  | 
|
| 25062 | 1785  | 
  assumes "finite N" and "N \<noteq> {}"
 | 
1786  | 
  shows "k + Max N = Max {k + m | m. m \<in> N}"
 | 
|
1787  | 
proof -  | 
|
1788  | 
have "\<And>x y. k + max x y = max (k + x) (k + y)"  | 
|
1789  | 
by (simp add: max_def not_le)  | 
|
1790  | 
(blast intro: antisym less_imp_le add_left_mono)  | 
|
1791  | 
with assms show ?thesis  | 
|
1792  | 
using hom_Max_commute [of "plus k" N]  | 
|
1793  | 
by simp (blast intro: arg_cong [where f = Max])  | 
|
1794  | 
qed  | 
|
| 22917 | 1795  | 
|
1796  | 
end  | 
|
1797  | 
||
| 35034 | 1798  | 
context linordered_ab_group_add  | 
1799  | 
begin  | 
|
1800  | 
||
1801  | 
lemma minus_Max_eq_Min [simp]:  | 
|
1802  | 
  "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - (Max S) = Min (uminus ` S)"
 | 
|
1803  | 
by (induct S rule: finite_ne_induct) (simp_all add: minus_max_eq_min)  | 
|
1804  | 
||
1805  | 
lemma minus_Min_eq_Max [simp]:  | 
|
1806  | 
  "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - (Min S) = Max (uminus ` S)"
 | 
|
1807  | 
by (induct S rule: finite_ne_induct) (simp_all add: minus_min_eq_max)  | 
|
1808  | 
||
1809  | 
end  | 
|
1810  | 
||
| 
25571
 
c9e39eafc7a0
instantiation target rather than legacy instance
 
haftmann 
parents: 
25502 
diff
changeset
 | 
1811  | 
end  |