| author | wenzelm | 
| Thu, 07 Nov 2024 13:22:59 +0100 | |
| changeset 81387 | c677755779f5 | 
| parent 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Algebra/Congruence.thy | 
| 35849 | 2 | Author: Clemens Ballarin, started 3 January 2008 | 
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 3 | with thanks to Paulo Emílio de Vilhena | 
| 27701 | 4 | *) | 
| 5 | ||
| 35849 | 6 | theory Congruence | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67999diff
changeset | 7 | imports | 
| 68188 
2af1f142f855
move FuncSet back to HOL-Library (amending 493b818e8e10)
 immler parents: 
68072diff
changeset | 8 | Main | 
| 
2af1f142f855
move FuncSet back to HOL-Library (amending 493b818e8e10)
 immler parents: 
68072diff
changeset | 9 | "HOL-Library.FuncSet" | 
| 35849 | 10 | begin | 
| 27701 | 11 | |
| 61382 | 12 | section \<open>Objects\<close> | 
| 27701 | 13 | |
| 61382 | 14 | subsection \<open>Structure with Carrier Set.\<close> | 
| 27701 | 15 | |
| 16 | record 'a partial_object = | |
| 17 | carrier :: "'a set" | |
| 18 | ||
| 65099 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 ballarin parents: 
63167diff
changeset | 19 | lemma funcset_carrier: | 
| 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 ballarin parents: 
63167diff
changeset | 20 | "\<lbrakk> f \<in> carrier X \<rightarrow> carrier Y; x \<in> carrier X \<rbrakk> \<Longrightarrow> f x \<in> carrier Y" | 
| 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 ballarin parents: 
63167diff
changeset | 21 | by (fact funcset_mem) | 
| 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 ballarin parents: 
63167diff
changeset | 22 | |
| 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 ballarin parents: 
63167diff
changeset | 23 | lemma funcset_carrier': | 
| 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 ballarin parents: 
63167diff
changeset | 24 | "\<lbrakk> f \<in> carrier A \<rightarrow> carrier A; x \<in> carrier A \<rbrakk> \<Longrightarrow> f x \<in> carrier A" | 
| 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 ballarin parents: 
63167diff
changeset | 25 | by (fact funcset_mem) | 
| 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 ballarin parents: 
63167diff
changeset | 26 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27701diff
changeset | 27 | |
| 63167 | 28 | subsection \<open>Structure with Carrier and Equivalence Relation \<open>eq\<close>\<close> | 
| 27701 | 29 | |
| 30 | record 'a eq_object = "'a partial_object" + | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69895diff
changeset | 31 | eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl \<open>.=\<index>\<close> 50) | 
| 27701 | 32 | |
| 35847 | 33 | definition | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69895diff
changeset | 34 | elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl \<open>.\<in>\<index>\<close> 50) | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 35 | where "x .\<in>\<^bsub>S\<^esub> A \<longleftrightarrow> (\<exists>y \<in> A. x .=\<^bsub>S\<^esub> y)" | 
| 27701 | 36 | |
| 35847 | 37 | definition | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69895diff
changeset | 38 |   set_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl \<open>{.=}\<index>\<close> 50)
 | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 39 |   where "A {.=}\<^bsub>S\<^esub> B \<longleftrightarrow> ((\<forall>x \<in> A. x .\<in>\<^bsub>S\<^esub> B) \<and> (\<forall>x \<in> B. x .\<in>\<^bsub>S\<^esub> A))"
 | 
| 27701 | 40 | |
| 35847 | 41 | definition | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69895diff
changeset | 42 | eq_class_of :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set" (\<open>class'_of\<index>\<close>) | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 43 |   where "class_of\<^bsub>S\<^esub> x = {y \<in> carrier S. x .=\<^bsub>S\<^esub> y}"
 | 
| 27701 | 44 | |
| 35847 | 45 | definition | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69895diff
changeset | 46 |   eq_classes :: "_ \<Rightarrow> ('a set) set" (\<open>classes\<index>\<close>)
 | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 47 |   where "classes\<^bsub>S\<^esub> = {class_of\<^bsub>S\<^esub> x | x. x \<in> carrier S}"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 48 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 49 | definition | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69895diff
changeset | 50 | eq_closure_of :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set" (\<open>closure'_of\<index>\<close>) | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 51 |   where "closure_of\<^bsub>S\<^esub> A = {y \<in> carrier S. y .\<in>\<^bsub>S\<^esub> A}"
 | 
| 27701 | 52 | |
| 35847 | 53 | definition | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69895diff
changeset | 54 | eq_is_closed :: "_ \<Rightarrow> 'a set \<Rightarrow> bool" (\<open>is'_closed\<index>\<close>) | 
| 35848 
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
 wenzelm parents: 
35847diff
changeset | 55 | where "is_closed\<^bsub>S\<^esub> A \<longleftrightarrow> A \<subseteq> carrier S \<and> closure_of\<^bsub>S\<^esub> A = A" | 
| 27701 | 56 | |
| 35355 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
29237diff
changeset | 57 | abbreviation | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69895diff
changeset | 58 | not_eq :: "_ \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl \<open>.\<noteq>\<index>\<close> 50) | 
| 67091 | 59 | where "x .\<noteq>\<^bsub>S\<^esub> y \<equiv> \<not>(x .=\<^bsub>S\<^esub> y)" | 
| 27701 | 60 | |
| 35355 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
29237diff
changeset | 61 | abbreviation | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69895diff
changeset | 62 | not_elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl \<open>.\<notin>\<index>\<close> 50) | 
| 67091 | 63 | where "x .\<notin>\<^bsub>S\<^esub> A \<equiv> \<not>(x .\<in>\<^bsub>S\<^esub> A)" | 
| 35355 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
29237diff
changeset | 64 | |
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
29237diff
changeset | 65 | abbreviation | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69895diff
changeset | 66 |   set_not_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl \<open>{.\<noteq>}\<index>\<close> 50)
 | 
| 67091 | 67 |   where "A {.\<noteq>}\<^bsub>S\<^esub> B \<equiv> \<not>(A {.=}\<^bsub>S\<^esub> B)"
 | 
| 27701 | 68 | |
| 69 | locale equivalence = | |
| 70 | fixes S (structure) | |
| 71 | assumes refl [simp, intro]: "x \<in> carrier S \<Longrightarrow> x .= x" | |
| 72 | and sym [sym]: "\<lbrakk> x .= y; x \<in> carrier S; y \<in> carrier S \<rbrakk> \<Longrightarrow> y .= x" | |
| 40293 | 73 | and trans [trans]: | 
| 74 | "\<lbrakk> x .= y; y .= z; x \<in> carrier S; y \<in> carrier S; z \<in> carrier S \<rbrakk> \<Longrightarrow> x .= z" | |
| 27701 | 75 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 76 | lemma equivalenceI: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 77 | fixes P :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and E :: "'a set" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 78 | assumes refl: "\<And>x. \<lbrakk> x \<in> E \<rbrakk> \<Longrightarrow> P x x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 79 | and sym: "\<And>x y. \<lbrakk> x \<in> E; y \<in> E \<rbrakk> \<Longrightarrow> P x y \<Longrightarrow> P y x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 80 | and trans: "\<And>x y z. \<lbrakk> x \<in> E; y \<in> E; z \<in> E \<rbrakk> \<Longrightarrow> P x y \<Longrightarrow> P y z \<Longrightarrow> P x z" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 81 | shows "equivalence \<lparr> carrier = E, eq = P \<rparr>" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 82 | unfolding equivalence_def using assms | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 83 | by (metis eq_object.select_convs(1) partial_object.select_convs(1)) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 84 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 85 | locale partition = | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 86 |   fixes A :: "'a set" and B :: "('a set) set"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 87 | assumes unique_class: "\<And>a. a \<in> A \<Longrightarrow> \<exists>!b \<in> B. a \<in> b" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 88 | and incl: "\<And>b. b \<in> B \<Longrightarrow> b \<subseteq> A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 89 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 90 | lemma equivalence_subset: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 91 | assumes "equivalence L" "A \<subseteq> carrier L" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 92 | shows "equivalence (L\<lparr> carrier := A \<rparr>)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 93 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 94 | interpret L: equivalence L | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 95 | by (simp add: assms) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 96 | show ?thesis | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 97 | by (unfold_locales, simp_all add: L.sym assms rev_subsetD, meson L.trans assms(2) contra_subsetD) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 98 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 99 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 100 | |
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27701diff
changeset | 101 | (* Lemmas by Stephan Hohe *) | 
| 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27701diff
changeset | 102 | |
| 27701 | 103 | lemma elemI: | 
| 104 | fixes R (structure) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 105 | assumes "a' \<in> A" "a .= a'" | 
| 27701 | 106 | shows "a .\<in> A" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 107 | unfolding elem_def using assms by auto | 
| 27701 | 108 | |
| 109 | lemma (in equivalence) elem_exact: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 110 | assumes "a \<in> carrier S" "a \<in> A" | 
| 27701 | 111 | shows "a .\<in> A" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 112 | unfolding elem_def using assms by auto | 
| 27701 | 113 | |
| 114 | lemma elemE: | |
| 115 | fixes S (structure) | |
| 116 | assumes "a .\<in> A" | |
| 117 | and "\<And>a'. \<lbrakk>a' \<in> A; a .= a'\<rbrakk> \<Longrightarrow> P" | |
| 118 | shows "P" | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 119 | using assms unfolding elem_def by auto | 
| 27701 | 120 | |
| 121 | lemma (in equivalence) elem_cong_l [trans]: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 122 | assumes "a \<in> carrier S" "a' \<in> carrier S" "A \<subseteq> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 123 | and "a' .= a" "a .\<in> A" | 
| 27701 | 124 | shows "a' .\<in> A" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 125 | using assms by (meson elem_def trans subsetCE) | 
| 27701 | 126 | |
| 127 | lemma (in equivalence) elem_subsetD: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 128 | assumes "A \<subseteq> B" "a .\<in> A" | 
| 27701 | 129 | shows "a .\<in> B" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 130 | using assms by (fast intro: elemI elim: elemE dest: subsetD) | 
| 27701 | 131 | |
| 132 | lemma (in equivalence) mem_imp_elem [simp, intro]: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 133 | assumes "x \<in> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 134 | shows "x \<in> A \<Longrightarrow> x .\<in> A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 135 | using assms unfolding elem_def by blast | 
| 27701 | 136 | |
| 137 | lemma set_eqI: | |
| 138 | fixes R (structure) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 139 | assumes "\<And>a. a \<in> A \<Longrightarrow> a .\<in> B" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 140 | and "\<And>b. b \<in> B \<Longrightarrow> b .\<in> A" | 
| 27701 | 141 |   shows "A {.=} B"
 | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 142 | using assms unfolding set_eq_def by auto | 
| 27701 | 143 | |
| 144 | lemma set_eqI2: | |
| 145 | fixes R (structure) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 146 | assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b \<in> B. a .= b" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 147 | and "\<And>b. b \<in> B \<Longrightarrow> \<exists>a \<in> A. b .= a" | 
| 27701 | 148 |   shows "A {.=} B"
 | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 149 | using assms by (simp add: set_eqI elem_def) | 
| 27701 | 150 | |
| 151 | lemma set_eqD1: | |
| 152 | fixes R (structure) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 153 |   assumes "A {.=} A'" and "a \<in> A"
 | 
| 27701 | 154 | shows "\<exists>a'\<in>A'. a .= a'" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 155 | using assms by (simp add: set_eq_def elem_def) | 
| 27701 | 156 | |
| 157 | lemma set_eqD2: | |
| 158 | fixes R (structure) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 159 |   assumes "A {.=} A'" and "a' \<in> A'"
 | 
| 27701 | 160 | shows "\<exists>a\<in>A. a' .= a" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 161 | using assms by (simp add: set_eq_def elem_def) | 
| 27701 | 162 | |
| 163 | lemma set_eqE: | |
| 164 | fixes R (structure) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 165 |   assumes "A {.=} B"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 166 | and "\<lbrakk> \<forall>a \<in> A. a .\<in> B; \<forall>b \<in> B. b .\<in> A \<rbrakk> \<Longrightarrow> P" | 
| 27701 | 167 | shows "P" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 168 | using assms unfolding set_eq_def by blast | 
| 27701 | 169 | |
| 170 | lemma set_eqE2: | |
| 171 | fixes R (structure) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 172 |   assumes "A {.=} B"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 173 | and "\<lbrakk> \<forall>a \<in> A. \<exists>b \<in> B. a .= b; \<forall>b \<in> B. \<exists>a \<in> A. b .= a \<rbrakk> \<Longrightarrow> P" | 
| 27701 | 174 | shows "P" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 175 | using assms unfolding set_eq_def by (simp add: elem_def) | 
| 27701 | 176 | |
| 177 | lemma set_eqE': | |
| 178 | fixes R (structure) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 179 |   assumes "A {.=} B" "a \<in> A" "b \<in> B"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 180 | and "\<And>a' b'. \<lbrakk> a' \<in> A; b' \<in> B \<rbrakk> \<Longrightarrow> b .= a' \<Longrightarrow> a .= b' \<Longrightarrow> P" | 
| 27701 | 181 | shows "P" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 182 | using assms by (meson set_eqE2) | 
| 27701 | 183 | |
| 184 | lemma (in equivalence) eq_elem_cong_r [trans]: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 185 |   assumes "A \<subseteq> carrier S" "A' \<subseteq> carrier S" "A {.=} A'"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 186 | shows "\<lbrakk> a \<in> carrier S \<rbrakk> \<Longrightarrow> a .\<in> A \<Longrightarrow> a .\<in> A'" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 187 | using assms by (meson elemE elem_cong_l set_eqE subset_eq) | 
| 27701 | 188 | |
| 189 | lemma (in equivalence) set_eq_sym [sym]: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 190 | assumes "A \<subseteq> carrier S" "B \<subseteq> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 191 |   shows "A {.=} B \<Longrightarrow> B {.=} A"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 192 | using assms unfolding set_eq_def elem_def by auto | 
| 27701 | 193 | |
| 194 | lemma (in equivalence) equal_set_eq_trans [trans]: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 195 |   "\<lbrakk> A = B; B {.=} C \<rbrakk> \<Longrightarrow> A {.=} C"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 196 | by simp | 
| 27701 | 197 | |
| 198 | lemma (in equivalence) set_eq_equal_trans [trans]: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 199 |   "\<lbrakk> A {.=} B; B = C \<rbrakk> \<Longrightarrow> A {.=} C"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 200 | by simp | 
| 27701 | 201 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 202 | lemma (in equivalence) set_eq_trans_aux: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 203 | assumes "A \<subseteq> carrier S" "B \<subseteq> carrier S" "C \<subseteq> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 204 |     and "A {.=} B" "B {.=} C"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 205 | shows "\<And>a. a \<in> A \<Longrightarrow> a .\<in> C" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 206 | using assms by (simp add: eq_elem_cong_r subset_iff) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 207 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 208 | corollary (in equivalence) set_eq_trans [trans]: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 209 | assumes "A \<subseteq> carrier S" "B \<subseteq> carrier S" "C \<subseteq> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 210 |     and "A {.=} B" " B {.=} C"
 | 
| 27701 | 211 |   shows "A {.=} C"
 | 
| 212 | proof (intro set_eqI) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 213 | show "\<And>a. a \<in> A \<Longrightarrow> a .\<in> C" using set_eq_trans_aux assms by blast | 
| 27701 | 214 | next | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 215 | show "\<And>b. b \<in> C \<Longrightarrow> b .\<in> A" using set_eq_trans_aux set_eq_sym assms by blast | 
| 27701 | 216 | qed | 
| 217 | ||
| 218 | lemma (in equivalence) is_closedI: | |
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 219 | assumes closed: "\<And>x y. \<lbrakk>x .= y; x \<in> A; y \<in> carrier S\<rbrakk> \<Longrightarrow> y \<in> A" | 
| 27701 | 220 | and S: "A \<subseteq> carrier S" | 
| 221 | shows "is_closed A" | |
| 222 | unfolding eq_is_closed_def eq_closure_of_def elem_def | |
| 223 | using S | |
| 224 | by (blast dest: closed sym) | |
| 225 | ||
| 226 | lemma (in equivalence) closure_of_eq: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 227 | assumes "A \<subseteq> carrier S" "x \<in> closure_of A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 228 | shows "\<lbrakk> x' \<in> carrier S; x .= x' \<rbrakk> \<Longrightarrow> x' \<in> closure_of A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 229 | using assms elem_cong_l sym unfolding eq_closure_of_def by blast | 
| 27701 | 230 | |
| 231 | lemma (in equivalence) is_closed_eq [dest]: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 232 | assumes "is_closed A" "x \<in> A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 233 | shows "\<lbrakk> x .= x'; x' \<in> carrier S \<rbrakk> \<Longrightarrow> x' \<in> A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 234 | using assms closure_of_eq [where A = A] unfolding eq_is_closed_def by simp | 
| 27701 | 235 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 236 | corollary (in equivalence) is_closed_eq_rev [dest]: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 237 | assumes "is_closed A" "x' \<in> A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 238 | shows "\<lbrakk> x .= x'; x \<in> carrier S \<rbrakk> \<Longrightarrow> x \<in> A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 239 | using sym is_closed_eq assms by (meson contra_subsetD eq_is_closed_def) | 
| 27701 | 240 | |
| 241 | lemma closure_of_closed [simp, intro]: | |
| 242 | fixes S (structure) | |
| 243 | shows "closure_of A \<subseteq> carrier S" | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 244 | unfolding eq_closure_of_def by auto | 
| 27701 | 245 | |
| 246 | lemma closure_of_memI: | |
| 247 | fixes S (structure) | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 248 | assumes "a .\<in> A" "a \<in> carrier S" | 
| 27701 | 249 | shows "a \<in> closure_of A" | 
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 250 | by (simp add: assms eq_closure_of_def) | 
| 27701 | 251 | |
| 252 | lemma closure_ofI2: | |
| 253 | fixes S (structure) | |
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 254 | assumes "a .= a'" and "a' \<in> A" and "a \<in> carrier S" | 
| 27701 | 255 | shows "a \<in> closure_of A" | 
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 256 | by (meson assms closure_of_memI elem_def) | 
| 27701 | 257 | |
| 258 | lemma closure_of_memE: | |
| 259 | fixes S (structure) | |
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 260 | assumes "a \<in> closure_of A" | 
| 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 261 | and "\<lbrakk>a \<in> carrier S; a .\<in> A\<rbrakk> \<Longrightarrow> P" | 
| 27701 | 262 | shows "P" | 
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 263 | using eq_closure_of_def assms by fastforce | 
| 27701 | 264 | |
| 265 | lemma closure_ofE2: | |
| 266 | fixes S (structure) | |
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 267 | assumes "a \<in> closure_of A" | 
| 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 268 | and "\<And>a'. \<lbrakk>a \<in> carrier S; a' \<in> A; a .= a'\<rbrakk> \<Longrightarrow> P" | 
| 27701 | 269 | shows "P" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 270 | using assms by (meson closure_of_memE elemE) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 271 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 272 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
68443diff
changeset | 273 | lemma (in partition) equivalence_from_partition: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 274 | "equivalence \<lparr> carrier = A, eq = (\<lambda>x y. y \<in> (THE b. b \<in> B \<and> x \<in> b))\<rparr>" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 275 | unfolding partition_def equivalence_def | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 276 | proof (auto) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 277 | let ?f = "\<lambda>x. THE b. b \<in> B \<and> x \<in> b" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 278 | show "\<And>x. x \<in> A \<Longrightarrow> x \<in> ?f x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 279 | using unique_class by (metis (mono_tags, lifting) theI') | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 280 | show "\<And>x y. \<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> y \<in> ?f x \<Longrightarrow> x \<in> ?f y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 281 | using unique_class by (metis (mono_tags, lifting) the_equality) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 282 | show "\<And>x y z. \<lbrakk> x \<in> A; y \<in> A; z \<in> A \<rbrakk> \<Longrightarrow> y \<in> ?f x \<Longrightarrow> z \<in> ?f y \<Longrightarrow> z \<in> ?f x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 283 | using unique_class by (metis (mono_tags, lifting) the_equality) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 284 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 285 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
68443diff
changeset | 286 | lemma (in partition) partition_coverture: "\<Union>B = A" \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 287 | by (meson Sup_le_iff UnionI unique_class incl subsetI subset_antisym) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 288 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
68443diff
changeset | 289 | lemma (in partition) disjoint_union: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 290 | assumes "b1 \<in> B" "b2 \<in> B" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 291 |     and "b1 \<inter> b2 \<noteq> {}"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 292 | shows "b1 = b2" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 293 | proof (rule ccontr) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 294 | assume "b1 \<noteq> b2" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 295 | obtain a where "a \<in> A" "a \<in> b1" "a \<in> b2" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 296 | using assms(2-3) incl by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 297 | thus False using unique_class \<open>b1 \<noteq> b2\<close> assms(1) assms(2) by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 298 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 299 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
68443diff
changeset | 300 | lemma partitionI: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 301 |   fixes A :: "'a set" and B :: "('a set) set"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 302 | assumes "\<Union>B = A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 303 |     and "\<And>b1 b2. \<lbrakk> b1 \<in> B; b2 \<in> B \<rbrakk> \<Longrightarrow> b1 \<inter> b2 \<noteq> {} \<Longrightarrow> b1 = b2"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 304 | shows "partition A B" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 305 | proof | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 306 | show "\<And>a. a \<in> A \<Longrightarrow> \<exists>!b. b \<in> B \<and> a \<in> b" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 307 | proof (rule ccontr) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 308 | fix a assume "a \<in> A" "\<nexists>!b. b \<in> B \<and> a \<in> b" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 309 | then obtain b1 b2 where "b1 \<in> B" "a \<in> b1" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 310 | and "b2 \<in> B" "a \<in> b2" "b1 \<noteq> b2" using assms(1) by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 311 | thus False using assms(2) by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 312 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 313 | next | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 314 | show "\<And>b. b \<in> B \<Longrightarrow> b \<subseteq> A" using assms(1) by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 315 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 316 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
68443diff
changeset | 317 | lemma (in partition) remove_elem: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 318 | assumes "b \<in> B" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 319 |   shows "partition (A - b) (B - {b})"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 320 | proof | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 321 |   show "\<And>a. a \<in> A - b \<Longrightarrow> \<exists>!b'. b' \<in> B - {b} \<and> a \<in> b'"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 322 | using unique_class by fastforce | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 323 | next | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 324 |   show "\<And>b'. b' \<in> B - {b} \<Longrightarrow> b' \<subseteq> A - b"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 325 | using assms unique_class incl partition_axioms partition_coverture by fastforce | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 326 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 327 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
68443diff
changeset | 328 | lemma disjoint_sum: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 329 | "\<lbrakk> finite B; finite A; partition A B\<rbrakk> \<Longrightarrow> (\<Sum>b\<in>B. \<Sum>a\<in>b. f a) = (\<Sum>a\<in>A. f a)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 330 | proof (induct arbitrary: A set: finite) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 331 | case empty thus ?case using partition.unique_class by fastforce | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 332 | next | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 333 | case (insert b B') | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 334 | have "(\<Sum>b\<in>(insert b B'). \<Sum>a\<in>b. f a) = (\<Sum>a\<in>b. f a) + (\<Sum>b\<in>B'. \<Sum>a\<in>b. f a)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 335 | by (simp add: insert.hyps(1) insert.hyps(2)) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 336 | also have "... = (\<Sum>a\<in>b. f a) + (\<Sum>a\<in>(A - b). f a)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 337 | using partition.remove_elem[of A "insert b B'" b] insert.hyps insert.prems | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 338 | by (metis Diff_insert_absorb finite_Diff insert_iff) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 339 | finally show "(\<Sum>b\<in>(insert b B'). \<Sum>a\<in>b. f a) = (\<Sum>a\<in>A. f a)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 340 | using partition.remove_elem[of A "insert b B'" b] insert.prems | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 341 | by (metis add.commute insert_iff partition.incl sum.subset_diff) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 342 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 343 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
68443diff
changeset | 344 | lemma (in partition) disjoint_sum: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 345 | assumes "finite A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 346 | shows "(\<Sum>b\<in>B. \<Sum>a\<in>b. f a) = (\<Sum>a\<in>A. f a)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 347 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 348 | have "finite B" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 349 | by (simp add: assms finite_UnionD partition_coverture) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 350 | thus ?thesis using disjoint_sum assms partition_axioms by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 351 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 352 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
68443diff
changeset | 353 | lemma (in equivalence) set_eq_insert_aux: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 354 | assumes "A \<subseteq> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 355 | and "x \<in> carrier S" "x' \<in> carrier S" "x .= x'" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 356 | and "y \<in> insert x A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 357 | shows "y .\<in> insert x' A" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 358 | by (metis assms(1) assms(4) assms(5) contra_subsetD elemI elem_exact insert_iff) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 359 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
68443diff
changeset | 360 | corollary (in equivalence) set_eq_insert: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 361 | assumes "A \<subseteq> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 362 | and "x \<in> carrier S" "x' \<in> carrier S" "x .= x'" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 363 |   shows "insert x A {.=} insert x' A"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 364 | by (meson set_eqI assms set_eq_insert_aux sym equivalence_axioms) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 365 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
68443diff
changeset | 366 | lemma (in equivalence) set_eq_pairI: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 367 | assumes xx': "x .= x'" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 368 | and carr: "x \<in> carrier S" "x' \<in> carrier S" "y \<in> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 369 |   shows "{x, y} {.=} {x', y}"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 370 | using assms set_eq_insert by simp | 
| 27701 | 371 | |
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 372 | lemma (in equivalence) closure_inclusion: | 
| 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 373 | assumes "A \<subseteq> B" | 
| 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 374 | shows "closure_of A \<subseteq> closure_of B" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 375 | unfolding eq_closure_of_def using assms elem_subsetD by auto | 
| 27701 | 376 | |
| 377 | lemma (in equivalence) classes_small: | |
| 378 | assumes "is_closed B" | |
| 379 | and "A \<subseteq> B" | |
| 380 | shows "closure_of A \<subseteq> B" | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 381 | by (metis assms closure_inclusion eq_is_closed_def) | 
| 27701 | 382 | |
| 383 | lemma (in equivalence) classes_eq: | |
| 384 | assumes "A \<subseteq> carrier S" | |
| 385 |   shows "A {.=} closure_of A"
 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 386 | using assms by (blast intro: set_eqI elem_exact closure_of_memI elim: closure_of_memE) | 
| 27701 | 387 | |
| 388 | lemma (in equivalence) complete_classes: | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 389 | assumes "is_closed A" | 
| 27701 | 390 | shows "A = closure_of A" | 
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 391 | using assms by (simp add: eq_is_closed_def) | 
| 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 392 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 393 | lemma (in equivalence) closure_idem_weak: | 
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 394 |   "closure_of (closure_of A) {.=} closure_of A"
 | 
| 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 395 | by (simp add: classes_eq set_eq_sym) | 
| 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 396 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 397 | lemma (in equivalence) closure_idem_strong: | 
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 398 | assumes "A \<subseteq> carrier S" | 
| 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 399 | shows "closure_of (closure_of A) = closure_of A" | 
| 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 400 | using assms closure_of_eq complete_classes is_closedI by auto | 
| 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 401 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 402 | lemma (in equivalence) classes_consistent: | 
| 67999 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 403 | assumes "A \<subseteq> carrier S" | 
| 
1b05f74f2e5f
tidying up including contributions from Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
67091diff
changeset | 404 | shows "is_closed (closure_of A)" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 405 | using closure_idem_strong by (simp add: assms eq_is_closed_def) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 406 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 407 | lemma (in equivalence) classes_coverture: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 408 | "\<Union>classes = carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 409 | proof | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 410 | show "\<Union>classes \<subseteq> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 411 | unfolding eq_classes_def eq_class_of_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 412 | next | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 413 | show "carrier S \<subseteq> \<Union>classes" unfolding eq_classes_def eq_class_of_def | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 414 | proof | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 415 | fix x assume "x \<in> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 416 |     hence "x \<in> {y \<in> carrier S. x .= y}" using refl by simp
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 417 |     thus "x \<in> \<Union>{{y \<in> carrier S. x .= y} |x. x \<in> carrier S}" by blast
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 418 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 419 | qed | 
| 27701 | 420 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 421 | lemma (in equivalence) disjoint_union: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 422 | assumes "class1 \<in> classes" "class2 \<in> classes" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 423 |     and "class1 \<inter> class2 \<noteq> {}"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 424 | shows "class1 = class2" | 
| 65099 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 ballarin parents: 
63167diff
changeset | 425 | proof - | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 426 | obtain x y where x: "x \<in> carrier S" "class1 = class_of x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 427 | and y: "y \<in> carrier S" "class2 = class_of y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 428 | using assms(1-2) unfolding eq_classes_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 429 | obtain z where z: "z \<in> carrier S" "z \<in> class1 \<inter> class2" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 430 | using assms classes_coverture by fastforce | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 431 | hence "x .= z \<and> y .= z" using x y unfolding eq_class_of_def by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 432 | hence "x .= y" using x y z trans sym by meson | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 433 | hence "class_of x = class_of y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 434 | unfolding eq_class_of_def using local.sym trans x y by blast | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 435 | thus ?thesis using x y by simp | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 436 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 437 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 438 | lemma (in equivalence) partition_from_equivalence: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 439 | "partition (carrier S) classes" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 440 | proof (intro partitionI) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 441 | show "\<Union>classes = carrier S" using classes_coverture by simp | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 442 | next | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 443 | show "\<And>class1 class2. \<lbrakk> class1 \<in> classes; class2 \<in> classes \<rbrakk> \<Longrightarrow> | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 444 |                           class1 \<inter> class2 \<noteq> {} \<Longrightarrow> class1 = class2"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 445 | using disjoint_union by simp | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 446 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 447 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 448 | lemma (in equivalence) disjoint_sum: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 449 | assumes "finite (carrier S)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 450 | shows "(\<Sum>c\<in>classes. \<Sum>x\<in>c. f x) = (\<Sum>x\<in>(carrier S). f x)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 451 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 452 | have "finite classes" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 453 | unfolding eq_classes_def using assms by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68188diff
changeset | 454 | thus ?thesis using disjoint_sum assms partition_from_equivalence by blast | 
| 65099 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 ballarin parents: 
63167diff
changeset | 455 | qed | 
| 
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
 ballarin parents: 
63167diff
changeset | 456 | |
| 27701 | 457 | end |