src/HOL/Algebra/Congruence.thy
author wenzelm
Sun, 10 Mar 2019 23:23:03 +0100
changeset 69895 6b03a8cf092d
parent 68443 43055b016688
permissions -rw-r--r--
more formal contributors (with the help of the history);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41959
b460124855b8 tuned headers;
wenzelm
parents: 40293
diff changeset
     1
(*  Title:      HOL/Algebra/Congruence.thy
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     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: 67091
diff changeset
     3
    with thanks to Paulo Emílio de Vilhena
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
     4
*)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
     5
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     6
theory Congruence
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67999
diff changeset
     7
  imports
68188
2af1f142f855 move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents: 68072
diff changeset
     8
    Main
2af1f142f855 move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents: 68072
diff changeset
     9
    "HOL-Library.FuncSet"
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
    10
begin
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    11
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 44471
diff changeset
    12
section \<open>Objects\<close>
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    13
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 44471
diff changeset
    14
subsection \<open>Structure with Carrier Set.\<close>
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    15
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    16
record 'a partial_object =
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    17
  carrier :: "'a set"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    18
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
    19
lemma funcset_carrier:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff 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: 63167
diff changeset
    21
  by (fact funcset_mem)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
    22
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
    23
lemma funcset_carrier':
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff 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: 63167
diff changeset
    25
  by (fact funcset_mem)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
    26
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27701
diff changeset
    27
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61382
diff changeset
    28
subsection \<open>Structure with Carrier and Equivalence Relation \<open>eq\<close>\<close>
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    29
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    30
record 'a eq_object = "'a partial_object" +
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    31
  eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl ".=\<index>" 50)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    32
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35355
diff changeset
    33
definition
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    34
  elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl ".\<in>\<index>" 50)
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
    35
  where "x .\<in>\<^bsub>S\<^esub> A \<longleftrightarrow> (\<exists>y \<in> A. x .=\<^bsub>S\<^esub> y)"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    36
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35355
diff changeset
    37
definition
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    38
  set_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "{.=}\<index>" 50)
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff 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
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    40
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35355
diff changeset
    41
definition
44471
3c2b2c4a7c1c tuned syntax -- avoid ambiguities;
wenzelm
parents: 41959
diff changeset
    42
  eq_class_of :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set" ("class'_of\<index>")
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
    43
  where "class_of\<^bsub>S\<^esub> x = {y \<in> carrier S. x .=\<^bsub>S\<^esub> y}"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    44
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35355
diff changeset
    45
definition
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
    46
  eq_classes :: "_ \<Rightarrow> ('a set) set" ("classes\<index>")
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
    48
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
    49
definition
44471
3c2b2c4a7c1c tuned syntax -- avoid ambiguities;
wenzelm
parents: 41959
diff changeset
    50
  eq_closure_of :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set" ("closure'_of\<index>")
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
    51
  where "closure_of\<^bsub>S\<^esub> A = {y \<in> carrier S. y .\<in>\<^bsub>S\<^esub> A}"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    52
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35355
diff changeset
    53
definition
44471
3c2b2c4a7c1c tuned syntax -- avoid ambiguities;
wenzelm
parents: 41959
diff changeset
    54
  eq_is_closed :: "_ \<Rightarrow> 'a set \<Rightarrow> bool" ("is'_closed\<index>")
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
    55
  where "is_closed\<^bsub>S\<^esub> A \<longleftrightarrow> A \<subseteq> carrier S \<and> closure_of\<^bsub>S\<^esub> A = A"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    56
35355
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 29237
diff changeset
    57
abbreviation
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    58
  not_eq :: "_ \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl ".\<noteq>\<index>" 50)
67091
1393c2340eec more symbols;
wenzelm
parents: 66453
diff changeset
    59
  where "x .\<noteq>\<^bsub>S\<^esub> y \<equiv> \<not>(x .=\<^bsub>S\<^esub> y)"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    60
35355
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 29237
diff changeset
    61
abbreviation
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 29237
diff changeset
    62
  not_elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl ".\<notin>\<index>" 50)
67091
1393c2340eec more symbols;
wenzelm
parents: 66453
diff changeset
    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: 29237
diff changeset
    64
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 29237
diff changeset
    65
abbreviation
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 29237
diff changeset
    66
  set_not_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "{.\<noteq>}\<index>" 50)
67091
1393c2340eec more symbols;
wenzelm
parents: 66453
diff changeset
    67
  where "A {.\<noteq>}\<^bsub>S\<^esub> B \<equiv> \<not>(A {.=}\<^bsub>S\<^esub> B)"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    68
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    69
locale equivalence =
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    70
  fixes S (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    71
  assumes refl [simp, intro]: "x \<in> carrier S \<Longrightarrow> x .= x"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    72
    and sym [sym]: "\<lbrakk> x .= y; x \<in> carrier S; y \<in> carrier S \<rbrakk> \<Longrightarrow> y .= x"
40293
cd932ab8cb59 Minor reformat.
ballarin
parents: 35849
diff changeset
    73
    and trans [trans]:
cd932ab8cb59 Minor reformat.
ballarin
parents: 35849
diff changeset
    74
      "\<lbrakk> x .= y; y .= z; x \<in> carrier S; y \<in> carrier S; z \<in> carrier S \<rbrakk> \<Longrightarrow> x .= z"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    75
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
    76
lemma equivalenceI:
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
    84
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
    85
locale partition =
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
    89
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
    90
lemma equivalence_subset:
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff changeset
    93
proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
    94
  interpret L: equivalence L
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
    95
    by (simp add: assms)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
    96
  show ?thesis
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
    98
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
    99
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   100
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27701
diff changeset
   101
(* Lemmas by Stephan Hohe *)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27701
diff changeset
   102
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   103
lemma elemI:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   104
  fixes R (structure)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   105
  assumes "a' \<in> A" "a .= a'"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   106
  shows "a .\<in> A"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   107
  unfolding elem_def using assms by auto
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   108
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff changeset
   110
  assumes "a \<in> carrier S" "a \<in> A"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   111
  shows "a .\<in> A"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   112
  unfolding elem_def using assms by auto
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   113
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   114
lemma elemE:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   115
  fixes S (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   116
  assumes "a .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   117
    and "\<And>a'. \<lbrakk>a' \<in> A; a .= a'\<rbrakk> \<Longrightarrow> P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   118
  shows "P"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   119
  using assms unfolding elem_def by auto
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   120
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff 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: 68188
diff changeset
   123
    and "a' .= a" "a .\<in> A"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   124
  shows "a' .\<in> A"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   125
  using assms by (meson elem_def trans subsetCE)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   126
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff changeset
   128
  assumes "A \<subseteq> B" "a .\<in> A"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   129
  shows "a .\<in> B"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   130
  using assms by (fast intro: elemI elim: elemE dest: subsetD)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   131
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   135
  using assms unfolding elem_def by blast
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   136
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   137
lemma set_eqI:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   138
  fixes R (structure)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   140
    and   "\<And>b. b \<in> B \<Longrightarrow> b .\<in> A"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   141
  shows "A {.=} B"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   142
  using assms unfolding set_eq_def by auto
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   143
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   144
lemma set_eqI2:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   145
  fixes R (structure)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   147
    and   "\<And>b. b \<in> B \<Longrightarrow> \<exists>a \<in> A. b .= a"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   148
  shows "A {.=} B"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   149
  using assms by (simp add: set_eqI elem_def)  
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   150
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   151
lemma set_eqD1:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   152
  fixes R (structure)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   153
  assumes "A {.=} A'" and "a \<in> A"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff changeset
   155
  using assms by (simp add: set_eq_def elem_def)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   156
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   157
lemma set_eqD2:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   158
  fixes R (structure)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   159
  assumes "A {.=} A'" and "a' \<in> A'"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff changeset
   161
  using assms by (simp add: set_eq_def elem_def)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   162
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   163
lemma set_eqE:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   164
  fixes R (structure)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   165
  assumes "A {.=} B"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   166
    and "\<lbrakk> \<forall>a \<in> A. a .\<in> B; \<forall>b \<in> B. b .\<in> A \<rbrakk> \<Longrightarrow> P"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   167
  shows "P"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   168
  using assms unfolding set_eq_def by blast
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   169
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   170
lemma set_eqE2:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   171
  fixes R (structure)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   172
  assumes "A {.=} B"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   174
  shows "P"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   175
  using assms unfolding set_eq_def by (simp add: elem_def) 
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   176
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   177
lemma set_eqE':
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   178
  fixes R (structure)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   180
    and "\<And>a' b'. \<lbrakk> a' \<in> A; b' \<in> B \<rbrakk> \<Longrightarrow> b .= a' \<Longrightarrow>  a .= b' \<Longrightarrow> P"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   181
  shows "P"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   182
  using assms by (meson set_eqE2)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   183
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   187
  using assms by (meson elemE elem_cong_l set_eqE subset_eq)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   188
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   192
  using assms unfolding set_eq_def elem_def by auto
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   193
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff 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: 68188
diff changeset
   196
  by simp
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   197
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff 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: 68188
diff changeset
   200
  by simp
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   201
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   207
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   210
    and "A {.=} B" " B {.=} C"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   211
  shows "A {.=} C"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   212
proof (intro set_eqI)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   213
  show "\<And>a. a \<in> A \<Longrightarrow> a .\<in> C" using set_eq_trans_aux assms by blast 
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   214
next
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   215
  show "\<And>b. b \<in> C \<Longrightarrow> b .\<in> A" using set_eq_trans_aux set_eq_sym assms by blast
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   216
qed
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   217
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   218
lemma (in equivalence) is_closedI:
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67091
diff changeset
   219
  assumes closed: "\<And>x y. \<lbrakk>x .= y; x \<in> A; y \<in> carrier S\<rbrakk> \<Longrightarrow> y \<in> A"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   220
    and S: "A \<subseteq> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   221
  shows "is_closed A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   222
  unfolding eq_is_closed_def eq_closure_of_def elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   223
  using S
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   224
  by (blast dest: closed sym)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   225
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   229
  using assms elem_cong_l sym unfolding eq_closure_of_def by blast 
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   230
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   234
  using assms closure_of_eq [where A = A] unfolding eq_is_closed_def by simp
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   235
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   239
  using sym is_closed_eq assms by (meson contra_subsetD eq_is_closed_def)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   240
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   241
lemma closure_of_closed [simp, intro]:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   242
  fixes S (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff changeset
   244
  unfolding eq_closure_of_def by auto
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   245
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   246
lemma closure_of_memI:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   247
  fixes S (structure)
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   248
  assumes "a .\<in> A" "a \<in> carrier S"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 67091
diff changeset
   250
  by (simp add: assms eq_closure_of_def)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   251
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   252
lemma closure_ofI2:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   253
  fixes S (structure)
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67091
diff changeset
   254
  assumes "a .= a'" and "a' \<in> A" and "a \<in> carrier S"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 67091
diff changeset
   256
  by (meson assms closure_of_memI elem_def)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   257
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   258
lemma closure_of_memE:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   259
  fixes S (structure)
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67091
diff 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: 67091
diff changeset
   261
    and "\<lbrakk>a \<in> carrier S; a .\<in> A\<rbrakk> \<Longrightarrow> P"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   262
  shows "P"
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67091
diff changeset
   263
  using eq_closure_of_def assms by fastforce
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   264
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   265
lemma closure_ofE2:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   266
  fixes S (structure)
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67091
diff 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: 67091
diff changeset
   268
    and "\<And>a'. \<lbrakk>a \<in> carrier S; a' \<in> A; a .= a'\<rbrakk> \<Longrightarrow> P"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   269
  shows "P"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   271
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   272
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 68443
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   276
proof (auto)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   284
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   285
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 68443
diff 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: 68188
diff 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: 68188
diff changeset
   288
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 68443
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   292
  shows "b1 = b2"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   293
proof (rule ccontr)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   294
  assume "b1 \<noteq> b2"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   298
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   299
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 68443
diff 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: 68188
diff 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: 68188
diff changeset
   302
  assumes "\<Union>B = A"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   304
  shows "partition A B"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   305
proof
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   307
  proof (rule ccontr)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   312
  qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   313
next
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   315
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   316
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 68443
diff 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: 68188
diff changeset
   318
  assumes "b \<in> B"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   320
proof
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   323
next
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   326
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   327
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 68443
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   332
next
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   333
  case (insert b B')
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   342
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   343
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 68443
diff 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: 68188
diff changeset
   345
  assumes "finite A"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   347
proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   348
  have "finite B"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   351
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   352
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 68443
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   359
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 68443
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   365
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 68443
diff 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: 68188
diff changeset
   367
  assumes xx': "x .= x'"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   370
  using assms set_eq_insert by simp
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   371
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67091
diff changeset
   372
lemma (in equivalence) closure_inclusion:
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67091
diff changeset
   373
  assumes "A \<subseteq> B"
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67091
diff 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: 68188
diff changeset
   375
  unfolding eq_closure_of_def using assms elem_subsetD by auto
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   376
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   377
lemma (in equivalence) classes_small:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   378
  assumes "is_closed B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   379
    and "A \<subseteq> B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff changeset
   381
  by (metis assms closure_inclusion eq_is_closed_def)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   382
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   383
lemma (in equivalence) classes_eq:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   384
  assumes "A \<subseteq> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff changeset
   386
  using assms by (blast intro: set_eqI elem_exact closure_of_memI elim: closure_of_memE)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   387
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   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: 68188
diff changeset
   389
  assumes "is_closed A"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   390
  shows "A = closure_of A"
67999
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67091
diff 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: 67091
diff changeset
   392
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 67091
diff 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: 67091
diff 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: 67091
diff changeset
   396
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 67091
diff changeset
   398
  assumes "A \<subseteq> carrier S"
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67091
diff 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: 67091
diff 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: 67091
diff changeset
   401
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 67091
diff changeset
   403
  assumes "A \<subseteq> carrier S"
1b05f74f2e5f tidying up including contributions from Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67091
diff 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: 68188
diff 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: 68188
diff changeset
   406
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   408
  "\<Union>classes = carrier S"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   409
proof
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   412
next
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   414
  proof
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   418
  qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   419
qed
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   420
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   424
    shows "class1 = class2"
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   425
proof -
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   436
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   437
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   439
  "partition (carrier S) classes"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   440
proof (intro partitionI)
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   442
next
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 68188
diff 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: 68188
diff changeset
   446
qed
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   447
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   449
  assumes "finite (carrier S)"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff changeset
   451
proof -
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff changeset
   452
  have "finite classes"
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68188
diff 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: 68188
diff 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: 63167
diff changeset
   455
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   456
  
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   457
end