| author | wenzelm | 
| Sat, 30 May 2015 23:30:54 +0200 | |
| changeset 60318 | ab785001b51d | 
| parent 59000 | 6eb0725503fc | 
| child 60758 | d8d85a8172b5 | 
| permissions | -rw-r--r-- | 
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1 | (* Title: HOL/Lattices_Big.thy | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 2 | Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 3 | with contributions by Jeremy Avigad | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 4 | *) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 5 | |
| 58889 | 6 | section {* Big infimum (minimum) and supremum (maximum) over finite (non-empty) sets *}
 | 
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 7 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 8 | theory Lattices_Big | 
| 55089 
181751ad852f
swapped dependencies of 'Finite_Set' and 'Option' (to move BNF up)
 blanchet parents: 
54868diff
changeset | 9 | imports Finite_Set Option | 
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 10 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 11 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 12 | subsection {* Generic lattice operations over a set *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 13 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 14 | no_notation times (infixl "*" 70) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 15 | no_notation Groups.one ("1")
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 16 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 17 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 18 | subsubsection {* Without neutral element *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 19 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 20 | locale semilattice_set = semilattice | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 21 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 22 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 23 | interpretation comp_fun_idem f | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 24 | by default (simp_all add: fun_eq_iff left_commute) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 25 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 26 | definition F :: "'a set \<Rightarrow> 'a" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 27 | where | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 28 | eq_fold': "F A = the (Finite_Set.fold (\<lambda>x y. Some (case y of None \<Rightarrow> x | Some z \<Rightarrow> f x z)) None A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 29 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 30 | lemma eq_fold: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 31 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 32 | shows "F (insert x A) = Finite_Set.fold f x A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 33 | proof (rule sym) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 34 | let ?f = "\<lambda>x y. Some (case y of None \<Rightarrow> x | Some z \<Rightarrow> f x z)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 35 | interpret comp_fun_idem "?f" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 36 | by default (simp_all add: fun_eq_iff commute left_commute split: option.split) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 37 | from assms show "Finite_Set.fold f x A = F (insert x A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 38 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 39 | case empty then show ?case by (simp add: eq_fold') | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 40 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 41 | case (insert y B) then show ?case by (simp add: insert_commute [of x] eq_fold') | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 42 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 43 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 44 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 45 | lemma singleton [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 46 |   "F {x} = x"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 47 | by (simp add: eq_fold) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 48 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 49 | lemma insert_not_elem: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 50 |   assumes "finite A" and "x \<notin> A" and "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 51 | shows "F (insert x A) = x * F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 52 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 53 |   from `A \<noteq> {}` obtain b where "b \<in> A" by blast
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 54 | then obtain B where *: "A = insert b B" "b \<notin> B" by (blast dest: mk_disjoint_insert) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 55 | with `finite A` and `x \<notin> A` | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 56 | have "finite (insert x B)" and "b \<notin> insert x B" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 57 | then have "F (insert b (insert x B)) = x * F (insert b B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 58 | by (simp add: eq_fold) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 59 | then show ?thesis by (simp add: * insert_commute) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 60 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 61 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 62 | lemma in_idem: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 63 | assumes "finite A" and "x \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 64 | shows "x * F A = F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 65 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 66 |   from assms have "A \<noteq> {}" by auto
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 67 | with `finite A` show ?thesis using `x \<in> A` | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 68 | by (induct A rule: finite_ne_induct) (auto simp add: ac_simps insert_not_elem) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 69 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 70 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 71 | lemma insert [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 72 |   assumes "finite A" and "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 73 | shows "F (insert x A) = x * F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 74 | using assms by (cases "x \<in> A") (simp_all add: insert_absorb in_idem insert_not_elem) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 75 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 76 | lemma union: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 77 |   assumes "finite A" "A \<noteq> {}" and "finite B" "B \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 78 | shows "F (A \<union> B) = F A * F B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 79 | using assms by (induct A rule: finite_ne_induct) (simp_all add: ac_simps) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 80 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 81 | lemma remove: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 82 | assumes "finite A" and "x \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 83 |   shows "F A = (if A - {x} = {} then x else x * F (A - {x}))"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 84 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 85 | from assms obtain B where "A = insert x B" and "x \<notin> B" by (blast dest: mk_disjoint_insert) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 86 | with assms show ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 87 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 88 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 89 | lemma insert_remove: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 90 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 91 |   shows "F (insert x A) = (if A - {x} = {} then x else x * F (A - {x}))"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 92 | using assms by (cases "x \<in> A") (simp_all add: insert_absorb remove) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 93 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 94 | lemma subset: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 95 |   assumes "finite A" "B \<noteq> {}" and "B \<subseteq> A"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 96 | shows "F B * F A = F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 97 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 98 |   from assms have "A \<noteq> {}" and "finite B" by (auto dest: finite_subset)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 99 | with assms show ?thesis by (simp add: union [symmetric] Un_absorb1) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 100 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 101 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 102 | lemma closed: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 103 |   assumes "finite A" "A \<noteq> {}" and elem: "\<And>x y. x * y \<in> {x, y}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 104 | shows "F A \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 105 | using `finite A` `A \<noteq> {}` proof (induct rule: finite_ne_induct)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 106 | case singleton then show ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 107 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 108 | case insert with elem show ?case by force | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 109 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 110 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 111 | lemma hom_commute: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 112 | assumes hom: "\<And>x y. h (x * y) = h x * h y" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 113 |   and N: "finite N" "N \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 114 | shows "h (F N) = F (h ` N)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 115 | using N proof (induct rule: finite_ne_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 116 | case singleton thus ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 117 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 118 | case (insert n N) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 119 | then have "h (F (insert n N)) = h (n * F N)" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 120 | also have "\<dots> = h n * h (F N)" by (rule hom) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 121 | also have "h (F N) = F (h ` N)" by (rule insert) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 122 | also have "h n * \<dots> = F (insert (h n) (h ` N))" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 123 | using insert by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 124 | also have "insert (h n) (h ` N) = h ` insert n N" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 125 | finally show ?case . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 126 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 127 | |
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56140diff
changeset | 128 | lemma infinite: "\<not> finite A \<Longrightarrow> F A = the None" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56140diff
changeset | 129 | unfolding eq_fold' by (cases "finite (UNIV::'a set)") (auto intro: finite_subset fold_infinite) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56140diff
changeset | 130 | |
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 131 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 132 | |
| 54745 | 133 | locale semilattice_order_set = binary?: semilattice_order + semilattice_set | 
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 134 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 135 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 136 | lemma bounded_iff: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 137 |   assumes "finite A" and "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 138 | shows "x \<preceq> F A \<longleftrightarrow> (\<forall>a\<in>A. x \<preceq> a)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 139 | using assms by (induct rule: finite_ne_induct) (simp_all add: bounded_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 140 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 141 | lemma boundedI: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 142 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 143 |   assumes "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 144 | assumes "\<And>a. a \<in> A \<Longrightarrow> x \<preceq> a" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 145 | shows "x \<preceq> F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 146 | using assms by (simp add: bounded_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 147 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 148 | lemma boundedE: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 149 |   assumes "finite A" and "A \<noteq> {}" and "x \<preceq> F A"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 150 | obtains "\<And>a. a \<in> A \<Longrightarrow> x \<preceq> a" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 151 | using assms by (simp add: bounded_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 152 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 153 | lemma coboundedI: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 154 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 155 | and "a \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 156 | shows "F A \<preceq> a" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 157 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 158 |   from assms have "A \<noteq> {}" by auto
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 159 |   from `finite A` `A \<noteq> {}` `a \<in> A` show ?thesis
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 160 | proof (induct rule: finite_ne_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 161 | case singleton thus ?case by (simp add: refl) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 162 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 163 | case (insert x B) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 164 | from insert have "a = x \<or> a \<in> B" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 165 | then show ?case using insert by (auto intro: coboundedI2) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 166 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 167 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 168 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 169 | lemma antimono: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 170 |   assumes "A \<subseteq> B" and "A \<noteq> {}" and "finite B"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 171 | shows "F B \<preceq> F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 172 | proof (cases "A = B") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 173 | case True then show ?thesis by (simp add: refl) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 174 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 175 | case False | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 176 | have B: "B = A \<union> (B - A)" using `A \<subseteq> B` by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 177 | then have "F B = F (A \<union> (B - A))" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 178 | also have "\<dots> = F A * F (B - A)" using False assms by (subst union) (auto intro: finite_subset) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 179 | also have "\<dots> \<preceq> F A" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 180 | finally show ?thesis . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 181 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 182 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 183 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 184 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 185 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 186 | subsubsection {* With neutral element *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 187 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 188 | locale semilattice_neutr_set = semilattice_neutr | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 189 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 190 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 191 | interpretation comp_fun_idem f | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 192 | by default (simp_all add: fun_eq_iff left_commute) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 193 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 194 | definition F :: "'a set \<Rightarrow> 'a" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 195 | where | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 196 | eq_fold: "F A = Finite_Set.fold f 1 A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 197 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 198 | lemma infinite [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 199 | "\<not> finite A \<Longrightarrow> F A = 1" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 200 | by (simp add: eq_fold) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 201 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 202 | lemma empty [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 203 |   "F {} = 1"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 204 | by (simp add: eq_fold) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 205 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 206 | lemma insert [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 207 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 208 | shows "F (insert x A) = x * F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 209 | using assms by (simp add: eq_fold) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 210 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 211 | lemma in_idem: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 212 | assumes "finite A" and "x \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 213 | shows "x * F A = F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 214 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 215 |   from assms have "A \<noteq> {}" by auto
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 216 | with `finite A` show ?thesis using `x \<in> A` | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 217 | by (induct A rule: finite_ne_induct) (auto simp add: ac_simps) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 218 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 219 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 220 | lemma union: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 221 | assumes "finite A" and "finite B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 222 | shows "F (A \<union> B) = F A * F B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 223 | using assms by (induct A) (simp_all add: ac_simps) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 224 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 225 | lemma remove: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 226 | assumes "finite A" and "x \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 227 |   shows "F A = x * F (A - {x})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 228 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 229 | from assms obtain B where "A = insert x B" and "x \<notin> B" by (blast dest: mk_disjoint_insert) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 230 | with assms show ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 231 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 232 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 233 | lemma insert_remove: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 234 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 235 |   shows "F (insert x A) = x * F (A - {x})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 236 | using assms by (cases "x \<in> A") (simp_all add: insert_absorb remove) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 237 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 238 | lemma subset: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 239 | assumes "finite A" and "B \<subseteq> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 240 | shows "F B * F A = F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 241 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 242 | from assms have "finite B" by (auto dest: finite_subset) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 243 | with assms show ?thesis by (simp add: union [symmetric] Un_absorb1) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 244 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 245 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 246 | lemma closed: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 247 |   assumes "finite A" "A \<noteq> {}" and elem: "\<And>x y. x * y \<in> {x, y}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 248 | shows "F A \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 249 | using `finite A` `A \<noteq> {}` proof (induct rule: finite_ne_induct)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 250 | case singleton then show ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 251 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 252 | case insert with elem show ?case by force | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 253 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 254 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 255 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 256 | |
| 54745 | 257 | locale semilattice_order_neutr_set = binary?: semilattice_neutr_order + semilattice_neutr_set | 
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 258 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 259 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 260 | lemma bounded_iff: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 261 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 262 | shows "x \<preceq> F A \<longleftrightarrow> (\<forall>a\<in>A. x \<preceq> a)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 263 | using assms by (induct A) (simp_all add: bounded_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 264 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 265 | lemma boundedI: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 266 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 267 | assumes "\<And>a. a \<in> A \<Longrightarrow> x \<preceq> a" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 268 | shows "x \<preceq> F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 269 | using assms by (simp add: bounded_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 270 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 271 | lemma boundedE: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 272 | assumes "finite A" and "x \<preceq> F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 273 | obtains "\<And>a. a \<in> A \<Longrightarrow> x \<preceq> a" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 274 | using assms by (simp add: bounded_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 275 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 276 | lemma coboundedI: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 277 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 278 | and "a \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 279 | shows "F A \<preceq> a" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 280 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 281 |   from assms have "A \<noteq> {}" by auto
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 282 |   from `finite A` `A \<noteq> {}` `a \<in> A` show ?thesis
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 283 | proof (induct rule: finite_ne_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 284 | case singleton thus ?case by (simp add: refl) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 285 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 286 | case (insert x B) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 287 | from insert have "a = x \<or> a \<in> B" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 288 | then show ?case using insert by (auto intro: coboundedI2) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 289 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 290 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 291 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 292 | lemma antimono: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 293 | assumes "A \<subseteq> B" and "finite B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 294 | shows "F B \<preceq> F A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 295 | proof (cases "A = B") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 296 | case True then show ?thesis by (simp add: refl) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 297 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 298 | case False | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 299 | have B: "B = A \<union> (B - A)" using `A \<subseteq> B` by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 300 | then have "F B = F (A \<union> (B - A))" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 301 | also have "\<dots> = F A * F (B - A)" using False assms by (subst union) (auto intro: finite_subset) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 302 | also have "\<dots> \<preceq> F A" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 303 | finally show ?thesis . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 304 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 305 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 306 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 307 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 308 | notation times (infixl "*" 70) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 309 | notation Groups.one ("1")
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 310 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 311 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 312 | subsection {* Lattice operations on finite sets *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 313 | |
| 54868 | 314 | context semilattice_inf | 
| 315 | begin | |
| 316 | ||
| 317 | definition Inf_fin :: "'a set \<Rightarrow> 'a" ("\<Sqinter>\<^sub>f\<^sub>i\<^sub>n_" [900] 900)
 | |
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 318 | where | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 319 | "Inf_fin = semilattice_set.F inf" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 320 | |
| 54868 | 321 | sublocale Inf_fin!: semilattice_order_set inf less_eq less | 
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 322 | where | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 323 | "semilattice_set.F inf = Inf_fin" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 324 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 325 | show "semilattice_order_set inf less_eq less" .. | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 326 | then interpret Inf_fin!: semilattice_order_set inf less_eq less . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 327 | from Inf_fin_def show "semilattice_set.F inf = Inf_fin" by rule | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 328 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 329 | |
| 54868 | 330 | end | 
| 331 | ||
| 332 | context semilattice_sup | |
| 333 | begin | |
| 334 | ||
| 335 | definition Sup_fin :: "'a set \<Rightarrow> 'a" ("\<Squnion>\<^sub>f\<^sub>i\<^sub>n_" [900] 900)
 | |
| 336 | where | |
| 337 | "Sup_fin = semilattice_set.F sup" | |
| 338 | ||
| 339 | sublocale Sup_fin!: semilattice_order_set sup greater_eq greater | |
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 340 | where | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 341 | "semilattice_set.F sup = Sup_fin" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 342 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 343 | show "semilattice_order_set sup greater_eq greater" .. | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 344 | then interpret Sup_fin!: semilattice_order_set sup greater_eq greater . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 345 | from Sup_fin_def show "semilattice_set.F sup = Sup_fin" by rule | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 346 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 347 | |
| 54868 | 348 | end | 
| 349 | ||
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 350 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 351 | subsection {* Infimum and Supremum over non-empty sets *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 352 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 353 | context lattice | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 354 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 355 | |
| 54745 | 356 | lemma Inf_fin_le_Sup_fin [simp]: | 
| 357 |   assumes "finite A" and "A \<noteq> {}"
 | |
| 358 | shows "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<le> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA" | |
| 359 | proof - | |
| 360 |   from `A \<noteq> {}` obtain a where "a \<in> A" by blast
 | |
| 361 | with `finite A` have "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<le> a" by (rule Inf_fin.coboundedI) | |
| 362 | moreover from `finite A` `a \<in> A` have "a \<le> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA" by (rule Sup_fin.coboundedI) | |
| 363 | ultimately show ?thesis by (rule order_trans) | |
| 364 | qed | |
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 365 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 366 | lemma sup_Inf_absorb [simp]: | 
| 54745 | 367 | "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> \<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<squnion> a = a" | 
| 368 | by (rule sup_absorb2) (rule Inf_fin.coboundedI) | |
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 369 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 370 | lemma inf_Sup_absorb [simp]: | 
| 54745 | 371 | "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> a \<sqinter> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA = a" | 
| 372 | by (rule inf_absorb1) (rule Sup_fin.coboundedI) | |
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 373 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 374 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 375 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 376 | context distrib_lattice | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 377 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 378 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 379 | lemma sup_Inf1_distrib: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 380 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 381 |     and "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 382 |   shows "sup x (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA) = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup x a|a. a \<in> A}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 383 | using assms by (simp add: image_def Inf_fin.hom_commute [where h="sup x", OF sup_inf_distrib1]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 384 | (rule arg_cong [where f="Inf_fin"], blast) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 385 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 386 | lemma sup_Inf2_distrib: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 387 |   assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 388 |   shows "sup (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB) = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup a b|a b. a \<in> A \<and> b \<in> B}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 389 | using A proof (induct rule: finite_ne_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 390 | case singleton then show ?case | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 391 | by (simp add: sup_Inf1_distrib [OF B]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 392 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 393 | case (insert x A) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 394 |   have finB: "finite {sup x b |b. b \<in> B}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 395 | by (rule finite_surj [where f = "sup x", OF B(1)], auto) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 396 |   have finAB: "finite {sup a b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 397 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 398 |     have "{sup a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {sup a b})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 399 | by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 400 | thus ?thesis by(simp add: insert(1) B(1)) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 401 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 402 |   have ne: "{sup a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 403 | have "sup (\<Sqinter>\<^sub>f\<^sub>i\<^sub>n(insert x A)) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB) = sup (inf x (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA)) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 404 | using insert by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 405 | also have "\<dots> = inf (sup x (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB)) (sup (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB))" by(rule sup_inf_distrib2) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 406 |   also have "\<dots> = inf (\<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup x b|b. b \<in> B}) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup a b|a b. a \<in> A \<and> b \<in> B})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 407 | using insert by(simp add:sup_Inf1_distrib[OF B]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 408 |   also have "\<dots> = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n({sup x b |b. b \<in> B} \<union> {sup a b |a b. a \<in> A \<and> b \<in> B})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 409 | (is "_ = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n?M") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 410 | using B insert | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 411 | by (simp add: Inf_fin.union [OF finB _ finAB ne]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 412 |   also have "?M = {sup a b |a b. a \<in> insert x A \<and> b \<in> B}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 413 | by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 414 | finally show ?case . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 415 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 416 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 417 | lemma inf_Sup1_distrib: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 418 |   assumes "finite A" and "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 419 |   shows "inf x (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA) = \<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf x a|a. a \<in> A}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 420 | using assms by (simp add: image_def Sup_fin.hom_commute [where h="inf x", OF inf_sup_distrib1]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 421 | (rule arg_cong [where f="Sup_fin"], blast) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 422 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 423 | lemma inf_Sup2_distrib: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 424 |   assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 425 |   shows "inf (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB) = \<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf a b|a b. a \<in> A \<and> b \<in> B}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 426 | using A proof (induct rule: finite_ne_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 427 | case singleton thus ?case | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 428 | by(simp add: inf_Sup1_distrib [OF B]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 429 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 430 | case (insert x A) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 431 |   have finB: "finite {inf x b |b. b \<in> B}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 432 | by(rule finite_surj[where f = "%b. inf x b", OF B(1)], auto) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 433 |   have finAB: "finite {inf a b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 434 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 435 |     have "{inf a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {inf a b})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 436 | by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 437 | thus ?thesis by(simp add: insert(1) B(1)) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 438 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 439 |   have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 440 | have "inf (\<Squnion>\<^sub>f\<^sub>i\<^sub>n(insert x A)) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB) = inf (sup x (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA)) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 441 | using insert by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 442 | also have "\<dots> = sup (inf x (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB)) (inf (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB))" by(rule inf_sup_distrib2) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 443 |   also have "\<dots> = sup (\<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf x b|b. b \<in> B}) (\<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf a b|a b. a \<in> A \<and> b \<in> B})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 444 | using insert by(simp add:inf_Sup1_distrib[OF B]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 445 |   also have "\<dots> = \<Squnion>\<^sub>f\<^sub>i\<^sub>n({inf x b |b. b \<in> B} \<union> {inf a b |a b. a \<in> A \<and> b \<in> B})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 446 | (is "_ = \<Squnion>\<^sub>f\<^sub>i\<^sub>n?M") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 447 | using B insert | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 448 | by (simp add: Sup_fin.union [OF finB _ finAB ne]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 449 |   also have "?M = {inf a b |a b. a \<in> insert x A \<and> b \<in> B}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 450 | by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 451 | finally show ?case . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 452 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 453 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 454 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 455 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 456 | context complete_lattice | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 457 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 458 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 459 | lemma Inf_fin_Inf: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 460 |   assumes "finite A" and "A \<noteq> {}"
 | 
| 54868 | 461 | shows "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA = \<Sqinter>A" | 
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 462 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 463 | from assms obtain b B where "A = insert b B" and "finite B" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 464 | then show ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 465 | by (simp add: Inf_fin.eq_fold inf_Inf_fold_inf inf.commute [of b]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 466 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 467 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 468 | lemma Sup_fin_Sup: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 469 |   assumes "finite A" and "A \<noteq> {}"
 | 
| 54868 | 470 | shows "\<Squnion>\<^sub>f\<^sub>i\<^sub>nA = \<Squnion>A" | 
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 471 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 472 | from assms obtain b B where "A = insert b B" and "finite B" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 473 | then show ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 474 | by (simp add: Sup_fin.eq_fold sup_Sup_fold_sup sup.commute [of b]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 475 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 476 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 477 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 478 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 479 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 480 | subsection {* Minimum and Maximum over non-empty sets *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 481 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 482 | context linorder | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 483 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 484 | |
| 54864 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 485 | definition Min :: "'a set \<Rightarrow> 'a" | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 486 | where | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 487 | "Min = semilattice_set.F min" | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 488 | |
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 489 | definition Max :: "'a set \<Rightarrow> 'a" | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 490 | where | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 491 | "Max = semilattice_set.F max" | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 492 | |
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 493 | sublocale Min!: semilattice_order_set min less_eq less | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 494 | + Max!: semilattice_order_set max greater_eq greater | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 495 | where | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 496 | "semilattice_set.F min = Min" | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 497 | and "semilattice_set.F max = Max" | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 498 | proof - | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 499 | show "semilattice_order_set min less_eq less" by default (auto simp add: min_def) | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 500 | then interpret Min!: semilattice_order_set min less_eq less . | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 501 | show "semilattice_order_set max greater_eq greater" by default (auto simp add: max_def) | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 502 | then interpret Max!: semilattice_order_set max greater_eq greater . | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 503 | from Min_def show "semilattice_set.F min = Min" by rule | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 504 | from Max_def show "semilattice_set.F max = Max" by rule | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 505 | qed | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 506 | |
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 507 | end | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 508 | |
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 509 | text {* An aside: @{const Min}/@{const Max} on linear orders as special case of @{const Inf_fin}/@{const Sup_fin} *}
 | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 510 | |
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 511 | lemma Inf_fin_Min: | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 512 |   "Inf_fin = (Min :: 'a::{semilattice_inf, linorder} set \<Rightarrow> 'a)"
 | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 513 | by (simp add: Inf_fin_def Min_def inf_min) | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 514 | |
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 515 | lemma Sup_fin_Max: | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 516 |   "Sup_fin = (Max :: 'a::{semilattice_sup, linorder} set \<Rightarrow> 'a)"
 | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 517 | by (simp add: Sup_fin_def Max_def sup_max) | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 518 | |
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 519 | context linorder | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 520 | begin | 
| 
a064732223ad
abolished slightly odd global lattice interpretation for min/max
 haftmann parents: 
54863diff
changeset | 521 | |
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 522 | lemma dual_min: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 523 | "ord.min greater_eq = max" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 524 | by (auto simp add: ord.min_def max_def fun_eq_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 525 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 526 | lemma dual_max: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 527 | "ord.max greater_eq = min" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 528 | by (auto simp add: ord.max_def min_def fun_eq_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 529 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 530 | lemma dual_Min: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 531 | "linorder.Min greater_eq = Max" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 532 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 533 | interpret dual!: linorder greater_eq greater by (fact dual_linorder) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 534 | show ?thesis by (simp add: dual.Min_def dual_min Max_def) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 535 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 536 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 537 | lemma dual_Max: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 538 | "linorder.Max greater_eq = Min" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 539 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 540 | interpret dual!: linorder greater_eq greater by (fact dual_linorder) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 541 | show ?thesis by (simp add: dual.Max_def dual_max Min_def) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 542 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 543 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 544 | lemmas Min_singleton = Min.singleton | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 545 | lemmas Max_singleton = Max.singleton | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 546 | lemmas Min_insert = Min.insert | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 547 | lemmas Max_insert = Max.insert | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 548 | lemmas Min_Un = Min.union | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 549 | lemmas Max_Un = Max.union | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 550 | lemmas hom_Min_commute = Min.hom_commute | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 551 | lemmas hom_Max_commute = Max.hom_commute | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 552 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 553 | lemma Min_in [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 554 |   assumes "finite A" and "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 555 | shows "Min A \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 556 | using assms by (auto simp add: min_def Min.closed) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 557 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 558 | lemma Max_in [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 559 |   assumes "finite A" and "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 560 | shows "Max A \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 561 | using assms by (auto simp add: max_def Max.closed) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 562 | |
| 58467 | 563 | lemma Min_insert2: | 
| 564 | assumes "finite A" and min: "\<And>b. b \<in> A \<Longrightarrow> a \<le> b" | |
| 565 | shows "Min (insert a A) = a" | |
| 566 | proof (cases "A = {}")
 | |
| 567 | case True then show ?thesis by simp | |
| 568 | next | |
| 569 | case False with `finite A` have "Min (insert a A) = min a (Min A)" | |
| 570 | by simp | |
| 571 |   moreover from `finite A` `A \<noteq> {}` min have "a \<le> Min A" by simp
 | |
| 572 | ultimately show ?thesis by (simp add: min.absorb1) | |
| 573 | qed | |
| 574 | ||
| 575 | lemma Max_insert2: | |
| 576 | assumes "finite A" and max: "\<And>b. b \<in> A \<Longrightarrow> b \<le> a" | |
| 577 | shows "Max (insert a A) = a" | |
| 578 | proof (cases "A = {}")
 | |
| 579 | case True then show ?thesis by simp | |
| 580 | next | |
| 581 | case False with `finite A` have "Max (insert a A) = max a (Max A)" | |
| 582 | by simp | |
| 583 |   moreover from `finite A` `A \<noteq> {}` max have "Max A \<le> a" by simp
 | |
| 584 | ultimately show ?thesis by (simp add: max.absorb1) | |
| 585 | qed | |
| 586 | ||
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 587 | lemma Min_le [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 588 | assumes "finite A" and "x \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 589 | shows "Min A \<le> x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 590 | using assms by (fact Min.coboundedI) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 591 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 592 | lemma Max_ge [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 593 | assumes "finite A" and "x \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 594 | shows "x \<le> Max A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 595 | using assms by (fact Max.coboundedI) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 596 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 597 | lemma Min_eqI: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 598 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 599 | assumes "\<And>y. y \<in> A \<Longrightarrow> y \<ge> x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 600 | and "x \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 601 | shows "Min A = x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 602 | proof (rule antisym) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 603 |   from `x \<in> A` have "A \<noteq> {}" by auto
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 604 | with assms show "Min A \<ge> x" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 605 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 606 | from assms show "x \<ge> Min A" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 607 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 608 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 609 | lemma Max_eqI: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 610 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 611 | assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 612 | and "x \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 613 | shows "Max A = x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 614 | proof (rule antisym) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 615 |   from `x \<in> A` have "A \<noteq> {}" by auto
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 616 | with assms show "Max A \<le> x" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 617 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 618 | from assms show "x \<le> Max A" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 619 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 620 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 621 | context | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 622 | fixes A :: "'a set" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 623 |   assumes fin_nonempty: "finite A" "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 624 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 625 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 626 | lemma Min_ge_iff [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 627 | "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 628 | using fin_nonempty by (fact Min.bounded_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 629 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 630 | lemma Max_le_iff [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 631 | "Max A \<le> x \<longleftrightarrow> (\<forall>a\<in>A. a \<le> x)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 632 | using fin_nonempty by (fact Max.bounded_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 633 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 634 | lemma Min_gr_iff [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 635 | "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 636 | using fin_nonempty by (induct rule: finite_ne_induct) simp_all | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 637 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 638 | lemma Max_less_iff [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 639 | "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 640 | using fin_nonempty by (induct rule: finite_ne_induct) simp_all | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 641 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 642 | lemma Min_le_iff: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 643 | "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 644 | using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_le_iff_disj) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 645 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 646 | lemma Max_ge_iff: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 647 | "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 648 | using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: le_max_iff_disj) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 649 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 650 | lemma Min_less_iff: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 651 | "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 652 | using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_less_iff_disj) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 653 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 654 | lemma Max_gr_iff: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 655 | "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 656 | using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: less_max_iff_disj) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 657 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 658 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 659 | |
| 57800 | 660 | lemma Max_eq_if: | 
| 661 | assumes "finite A" "finite B" "\<forall>a\<in>A. \<exists>b\<in>B. a \<le> b" "\<forall>b\<in>B. \<exists>a\<in>A. b \<le> a" | |
| 662 | shows "Max A = Max B" | |
| 663 | proof cases | |
| 664 |   assume "A = {}" thus ?thesis using assms by simp
 | |
| 665 | next | |
| 666 |   assume "A \<noteq> {}" thus ?thesis using assms
 | |
| 667 | by(blast intro: antisym Max_in Max_ge_iff[THEN iffD2]) | |
| 668 | qed | |
| 669 | ||
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 670 | lemma Min_antimono: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 671 |   assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 672 | shows "Min N \<le> Min M" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 673 | using assms by (fact Min.antimono) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 674 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 675 | lemma Max_mono: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 676 |   assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 677 | shows "Max M \<le> Max N" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 678 | using assms by (fact Max.antimono) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 679 | |
| 56140 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
55803diff
changeset | 680 | end | 
| 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
55803diff
changeset | 681 | |
| 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
55803diff
changeset | 682 | context linorder (* FIXME *) | 
| 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
55803diff
changeset | 683 | begin | 
| 
ed92ce2ac88e
just one cumulative Proof_Context.facts, with uniform retrieval (including PIDE markup, completion etc.);
 wenzelm parents: 
55803diff
changeset | 684 | |
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 685 | lemma mono_Min_commute: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 686 | assumes "mono f" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 687 |   assumes "finite A" and "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 688 | shows "f (Min A) = Min (f ` A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 689 | proof (rule linorder_class.Min_eqI [symmetric]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 690 | from `finite A` show "finite (f ` A)" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 691 | from assms show "f (Min A) \<in> f ` A" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 692 | fix x | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 693 | assume "x \<in> f ` A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 694 | then obtain y where "y \<in> A" and "x = f y" .. | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 695 | with assms have "Min A \<le> y" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 696 | with `mono f` have "f (Min A) \<le> f y" by (rule monoE) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 697 | with `x = f y` show "f (Min A) \<le> x" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 698 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 699 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 700 | lemma mono_Max_commute: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 701 | assumes "mono f" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 702 |   assumes "finite A" and "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 703 | shows "f (Max A) = Max (f ` A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 704 | proof (rule linorder_class.Max_eqI [symmetric]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 705 | from `finite A` show "finite (f ` A)" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 706 | from assms show "f (Max A) \<in> f ` A" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 707 | fix x | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 708 | assume "x \<in> f ` A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 709 | then obtain y where "y \<in> A" and "x = f y" .. | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 710 | with assms have "y \<le> Max A" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 711 | with `mono f` have "f y \<le> f (Max A)" by (rule monoE) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 712 | with `x = f y` show "x \<le> f (Max A)" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 713 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 714 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 715 | lemma finite_linorder_max_induct [consumes 1, case_names empty insert]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 716 | assumes fin: "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 717 |   and empty: "P {}" 
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 718 | and insert: "\<And>b A. finite A \<Longrightarrow> \<forall>a\<in>A. a < b \<Longrightarrow> P A \<Longrightarrow> P (insert b A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 719 | shows "P A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 720 | using fin empty insert | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 721 | proof (induct rule: finite_psubset_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 722 | case (psubset A) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 723 |   have IH: "\<And>B. \<lbrakk>B < A; P {}; (\<And>A b. \<lbrakk>finite A; \<forall>a\<in>A. a<b; P A\<rbrakk> \<Longrightarrow> P (insert b A))\<rbrakk> \<Longrightarrow> P B" by fact 
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 724 | have fin: "finite A" by fact | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 725 |   have empty: "P {}" by fact
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 726 | have step: "\<And>b A. \<lbrakk>finite A; \<forall>a\<in>A. a < b; P A\<rbrakk> \<Longrightarrow> P (insert b A)" by fact | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 727 | show "P A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 728 |   proof (cases "A = {}")
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 729 |     assume "A = {}" 
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 730 |     then show "P A" using `P {}` by simp
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 731 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 732 |     let ?B = "A - {Max A}" 
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 733 | let ?A = "insert (Max A) ?B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 734 | have "finite ?B" using `finite A` by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 735 |     assume "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 736 | with `finite A` have "Max A : A" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 737 | then have A: "?A = A" using insert_Diff_single insert_absorb by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 738 |     then have "P ?B" using `P {}` step IH [of ?B] by blast
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 739 | moreover | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 740 | have "\<forall>a\<in>?B. a < Max A" using Max_ge [OF `finite A`] by fastforce | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 741 | ultimately show "P A" using A insert_Diff_single step [OF `finite ?B`] by fastforce | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 742 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 743 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 744 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 745 | lemma finite_linorder_min_induct [consumes 1, case_names empty insert]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 746 |   "\<lbrakk>finite A; P {}; \<And>b A. \<lbrakk>finite A; \<forall>a\<in>A. b < a; P A\<rbrakk> \<Longrightarrow> P (insert b A)\<rbrakk> \<Longrightarrow> P A"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 747 | by (rule linorder.finite_linorder_max_induct [OF dual_linorder]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 748 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 749 | lemma Least_Min: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 750 |   assumes "finite {a. P a}" and "\<exists>a. P a"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 751 |   shows "(LEAST a. P a) = Min {a. P a}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 752 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 753 |   { fix A :: "'a set"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 754 |     assume A: "finite A" "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 755 | have "(LEAST a. a \<in> A) = Min A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 756 | using A proof (induct A rule: finite_ne_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 757 | case singleton show ?case by (rule Least_equality) simp_all | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 758 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 759 | case (insert a A) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 760 | have "(LEAST b. b = a \<or> b \<in> A) = min a (LEAST a. a \<in> A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 761 | by (auto intro!: Least_equality simp add: min_def not_le Min_le_iff insert.hyps dest!: less_imp_le) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 762 | with insert show ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 763 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 764 |   } from this [of "{a. P a}"] assms show ?thesis by simp
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 765 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 766 | |
| 59000 | 767 | lemma infinite_growing: | 
| 768 |   assumes "X \<noteq> {}"
 | |
| 769 | assumes *: "\<And>x. x \<in> X \<Longrightarrow> \<exists>y\<in>X. y > x" | |
| 770 | shows "\<not> finite X" | |
| 771 | proof | |
| 772 | assume "finite X" | |
| 773 |   with `X \<noteq> {}` have "Max X \<in> X" "\<forall>x\<in>X. x \<le> Max X"
 | |
| 774 | by auto | |
| 775 | with *[of "Max X"] show False | |
| 776 | by auto | |
| 777 | qed | |
| 778 | ||
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 779 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 780 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 781 | context linordered_ab_semigroup_add | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 782 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 783 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 784 | lemma add_Min_commute: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 785 | fixes k | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 786 |   assumes "finite N" and "N \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 787 |   shows "k + Min N = Min {k + m | m. m \<in> N}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 788 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 789 | have "\<And>x y. k + min x y = min (k + x) (k + y)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 790 | by (simp add: min_def not_le) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 791 | (blast intro: antisym less_imp_le add_left_mono) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 792 | with assms show ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 793 | using hom_Min_commute [of "plus k" N] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 794 | by simp (blast intro: arg_cong [where f = Min]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 795 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 796 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 797 | lemma add_Max_commute: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 798 | fixes k | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 799 |   assumes "finite N" and "N \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 800 |   shows "k + Max N = Max {k + m | m. m \<in> N}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 801 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 802 | have "\<And>x y. k + max x y = max (k + x) (k + y)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 803 | by (simp add: max_def not_le) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 804 | (blast intro: antisym less_imp_le add_left_mono) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 805 | with assms show ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 806 | using hom_Max_commute [of "plus k" N] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 807 | by simp (blast intro: arg_cong [where f = Max]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 808 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 809 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 810 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 811 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 812 | context linordered_ab_group_add | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 813 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 814 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 815 | lemma minus_Max_eq_Min [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 816 |   "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - Max S = Min (uminus ` S)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 817 | by (induct S rule: finite_ne_induct) (simp_all add: minus_max_eq_min) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 818 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 819 | lemma minus_Min_eq_Max [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 820 |   "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - Min S = Max (uminus ` S)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 821 | by (induct S rule: finite_ne_induct) (simp_all add: minus_min_eq_max) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 822 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 823 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 824 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 825 | context complete_linorder | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 826 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 827 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 828 | lemma Min_Inf: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 829 |   assumes "finite A" and "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 830 | shows "Min A = Inf A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 831 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 832 | from assms obtain b B where "A = insert b B" and "finite B" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 833 | then show ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 834 | by (simp add: Min.eq_fold complete_linorder_inf_min [symmetric] inf_Inf_fold_inf inf.commute [of b]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 835 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 836 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 837 | lemma Max_Sup: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 838 |   assumes "finite A" and "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 839 | shows "Max A = Sup A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 840 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 841 | from assms obtain b B where "A = insert b B" and "finite B" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 842 | then show ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 843 | by (simp add: Max.eq_fold complete_linorder_sup_max [symmetric] sup_Sup_fold_sup sup.commute [of b]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 844 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 845 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 846 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 847 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 848 | end |