src/HOL/Algebra/Congruence.thy
author haftmann
Thu, 23 Nov 2017 17:03:27 +0000
changeset 67087 733017b19de9
parent 66453 cc19f7ca2ed6
child 67091 1393c2340eec
permissions -rw-r--r--
generalized more lemmas
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
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     3
    Copyright:  Clemens Ballarin
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
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
     7
imports 
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
     8
  Main
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 65099
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
44471
3c2b2c4a7c1c tuned syntax -- avoid ambiguities;
wenzelm
parents: 41959
diff changeset
    46
  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
    47
  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
    48
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35355
diff changeset
    49
definition
44471
3c2b2c4a7c1c tuned syntax -- avoid ambiguities;
wenzelm
parents: 41959
diff changeset
    50
  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
    51
  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
    52
35355
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 29237
diff changeset
    53
abbreviation
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    54
  not_eq :: "_ \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl ".\<noteq>\<index>" 50)
35355
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 29237
diff changeset
    55
  where "x .\<noteq>\<^bsub>S\<^esub> y == ~(x .=\<^bsub>S\<^esub> y)"
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
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 29237
diff changeset
    58
  not_elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl ".\<notin>\<index>" 50)
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 29237
diff changeset
    59
  where "x .\<notin>\<^bsub>S\<^esub> A == ~(x .\<in>\<^bsub>S\<^esub> A)"
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 29237
diff changeset
    60
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
  set_not_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "{.\<noteq>}\<index>" 50)
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 29237
diff changeset
    63
  where "A {.\<noteq>}\<^bsub>S\<^esub> B == ~(A {.=}\<^bsub>S\<^esub> B)"
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    64
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    65
locale equivalence =
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    66
  fixes S (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    67
  assumes refl [simp, intro]: "x \<in> carrier S \<Longrightarrow> x .= x"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    68
    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
    69
    and trans [trans]:
cd932ab8cb59 Minor reformat.
ballarin
parents: 35849
diff changeset
    70
      "\<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
    71
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27701
diff changeset
    72
(* Lemmas by Stephan Hohe *)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27701
diff changeset
    73
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    74
lemma elemI:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    75
  fixes R (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    76
  assumes "a' \<in> A" and "a .= a'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    77
  shows "a .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    78
unfolding elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    79
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    80
by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    81
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    82
lemma (in equivalence) elem_exact:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    83
  assumes "a \<in> carrier S" and "a \<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    84
  shows "a .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    85
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    86
by (fast intro: elemI)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    87
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    88
lemma elemE:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    89
  fixes S (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    90
  assumes "a .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    91
    and "\<And>a'. \<lbrakk>a' \<in> A; a .= a'\<rbrakk> \<Longrightarrow> P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    92
  shows "P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    93
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    94
unfolding elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    95
by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    96
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    97
lemma (in equivalence) elem_cong_l [trans]:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    98
  assumes cong: "a' .= a"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
    99
    and a: "a .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   100
    and carr: "a \<in> carrier S"  "a' \<in> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   101
    and Acarr: "A \<subseteq> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   102
  shows "a' .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   103
using a
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   104
apply (elim elemE, intro elemI)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   105
proof assumption
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   106
  fix b
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   107
  assume bA: "b \<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   108
  note [simp] = carr bA[THEN subsetD[OF Acarr]]
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   109
  note cong
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   110
  also assume "a .= b"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   111
  finally show "a' .= b" by simp
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   112
qed
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   113
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   114
lemma (in equivalence) elem_subsetD:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   115
  assumes "A \<subseteq> B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   116
    and aA: "a .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   117
  shows "a .\<in> B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   118
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   119
by (fast intro: elemI elim: elemE dest: subsetD)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   120
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   121
lemma (in equivalence) mem_imp_elem [simp, intro]:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   122
  "[| x \<in> A; x \<in> carrier S |] ==> x .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   123
  unfolding elem_def by blast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   124
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   125
lemma set_eqI:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   126
  fixes R (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   127
  assumes ltr: "\<And>a. a \<in> A \<Longrightarrow> a .\<in> B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   128
    and rtl: "\<And>b. b \<in> B \<Longrightarrow> b .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   129
  shows "A {.=} B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   130
unfolding set_eq_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   131
by (fast intro: ltr rtl)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   132
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   133
lemma set_eqI2:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   134
  fixes R (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   135
  assumes ltr: "\<And>a b. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a .= b"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   136
    and rtl: "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. b .= a"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   137
  shows "A {.=} B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   138
  by (intro set_eqI, unfold elem_def) (fast intro: ltr rtl)+
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   139
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   140
lemma set_eqD1:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   141
  fixes R (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   142
  assumes AA': "A {.=} A'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   143
    and "a \<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   144
  shows "\<exists>a'\<in>A'. a .= a'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   145
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   146
unfolding set_eq_def elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   147
by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   148
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   149
lemma set_eqD2:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   150
  fixes R (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   151
  assumes AA': "A {.=} A'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   152
    and "a' \<in> A'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   153
  shows "\<exists>a\<in>A. a' .= a"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   154
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   155
unfolding set_eq_def elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   156
by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   157
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   158
lemma set_eqE:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   159
  fixes R (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   160
  assumes AB: "A {.=} B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   161
    and r: "\<lbrakk>\<forall>a\<in>A. a .\<in> B; \<forall>b\<in>B. b .\<in> A\<rbrakk> \<Longrightarrow> P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   162
  shows "P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   163
using AB
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   164
unfolding set_eq_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   165
by (blast dest: r)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   166
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   167
lemma set_eqE2:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   168
  fixes R (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   169
  assumes AB: "A {.=} B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   170
    and r: "\<lbrakk>\<forall>a\<in>A. (\<exists>b\<in>B. a .= b); \<forall>b\<in>B. (\<exists>a\<in>A. b .= a)\<rbrakk> \<Longrightarrow> P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   171
  shows "P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   172
using AB
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   173
unfolding set_eq_def elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   174
by (blast dest: r)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   175
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   176
lemma set_eqE':
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   177
  fixes R (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   178
  assumes AB: "A {.=} B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   179
    and aA: "a \<in> A" and bB: "b \<in> B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   180
    and r: "\<And>a' b'. \<lbrakk>a' \<in> A; b .= a'; b' \<in> B; a .= b'\<rbrakk> \<Longrightarrow> P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   181
  shows "P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   182
proof -
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   183
  from AB aA
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   184
      have "\<exists>b'\<in>B. a .= b'" by (rule set_eqD1)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   185
  from this obtain b'
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   186
      where b': "b' \<in> B" "a .= b'" by auto
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   187
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   188
  from AB bB
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   189
      have "\<exists>a'\<in>A. b .= a'" by (rule set_eqD2)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   190
  from this obtain a'
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   191
      where a': "a' \<in> A" "b .= a'" by auto
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   192
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   193
  from a' b'
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   194
      show "P" by (rule r)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   195
qed
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   196
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   197
lemma (in equivalence) eq_elem_cong_r [trans]:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   198
  assumes a: "a .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   199
    and cong: "A {.=} A'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   200
    and carr: "a \<in> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   201
    and Carr: "A \<subseteq> carrier S" "A' \<subseteq> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   202
  shows "a .\<in> A'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   203
using a cong
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   204
proof (elim elemE set_eqE)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   205
  fix b
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   206
  assume bA: "b \<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   207
     and inA': "\<forall>b\<in>A. b .\<in> A'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   208
  note [simp] = carr Carr Carr[THEN subsetD] bA
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   209
  assume "a .= b"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   210
  also from bA inA'
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   211
       have "b .\<in> A'" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   212
  finally
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   213
       show "a .\<in> A'" by simp
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   214
qed
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   215
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   216
lemma (in equivalence) set_eq_sym [sym]:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   217
  assumes "A {.=} B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   218
    and "A \<subseteq> carrier S" "B \<subseteq> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   219
  shows "B {.=} A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   220
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   221
unfolding set_eq_def elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   222
by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   223
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   224
(* FIXME: the following two required in Isabelle 2008, not Isabelle 2007 *)
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27701
diff changeset
   225
(* alternatively, could declare lemmas [trans] = ssubst [where 'a = "'a set"] *)
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   226
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   227
lemma (in equivalence) equal_set_eq_trans [trans]:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   228
  assumes AB: "A = B" and BC: "B {.=} C"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   229
  shows "A {.=} C"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   230
  using AB BC by simp
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   231
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   232
lemma (in equivalence) set_eq_equal_trans [trans]:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   233
  assumes AB: "A {.=} B" and BC: "B = C"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   234
  shows "A {.=} C"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   235
  using AB BC by simp
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   236
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27701
diff changeset
   237
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   238
lemma (in equivalence) set_eq_trans [trans]:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   239
  assumes AB: "A {.=} B" and BC: "B {.=} C"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   240
    and carr: "A \<subseteq> carrier S"  "B \<subseteq> carrier S"  "C \<subseteq> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   241
  shows "A {.=} C"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   242
proof (intro set_eqI)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   243
  fix a
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   244
  assume aA: "a \<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   245
  with carr have "a \<in> carrier S" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   246
  note [simp] = carr this
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   247
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   248
  from aA
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   249
       have "a .\<in> A" by (simp add: elem_exact)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   250
  also note AB
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   251
  also note BC
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   252
  finally
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   253
       show "a .\<in> C" by simp
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   254
next
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   255
  fix c
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   256
  assume cC: "c \<in> C"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   257
  with carr have "c \<in> carrier S" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   258
  note [simp] = carr this
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   259
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   260
  from cC
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   261
       have "c .\<in> C" by (simp add: elem_exact)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   262
  also note BC[symmetric]
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   263
  also note AB[symmetric]
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   264
  finally
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   265
       show "c .\<in> A" by simp
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   266
qed
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   267
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   268
(* FIXME: generalise for insert *)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   269
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   270
(*
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   271
lemma (in equivalence) set_eq_insert:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   272
  assumes x: "x .= x'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   273
    and carr: "x \<in> carrier S" "x' \<in> carrier S" "A \<subseteq> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   274
  shows "insert x A {.=} insert x' A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   275
  unfolding set_eq_def elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   276
apply rule
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   277
apply rule
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   278
apply (case_tac "xa = x")
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   279
using x apply fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   280
apply (subgoal_tac "xa \<in> A") prefer 2 apply fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   281
apply (rule_tac x=xa in bexI)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   282
using carr apply (rule_tac refl) apply auto [1]
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   283
apply safe
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   284
*)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   285
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   286
lemma (in equivalence) set_eq_pairI:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   287
  assumes xx': "x .= x'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   288
    and carr: "x \<in> carrier S" "x' \<in> carrier S" "y \<in> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   289
  shows "{x, y} {.=} {x', y}"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   290
unfolding set_eq_def elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   291
proof safe
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   292
  have "x' \<in> {x', y}" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   293
  with xx' show "\<exists>b\<in>{x', y}. x .= b" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   294
next
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   295
  have "y \<in> {x', y}" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   296
  with carr show "\<exists>b\<in>{x', y}. y .= b" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   297
next
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   298
  have "x \<in> {x, y}" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   299
  with xx'[symmetric] carr
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   300
  show "\<exists>a\<in>{x, y}. x' .= a" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   301
next
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   302
  have "y \<in> {x, y}" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   303
  with carr show "\<exists>a\<in>{x, y}. y .= a" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   304
qed
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   305
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   306
lemma (in equivalence) is_closedI:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   307
  assumes closed: "!!x y. [| x .= y; x \<in> A; y \<in> carrier S |] ==> y \<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   308
    and S: "A \<subseteq> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   309
  shows "is_closed A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   310
  unfolding eq_is_closed_def eq_closure_of_def elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   311
  using S
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   312
  by (blast dest: closed sym)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   313
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   314
lemma (in equivalence) closure_of_eq:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   315
  "[| x .= x'; A \<subseteq> carrier S; x \<in> closure_of A; x \<in> carrier S; x' \<in> carrier S |] ==> x' \<in> closure_of A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   316
  unfolding eq_closure_of_def elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   317
  by (blast intro: trans sym)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   318
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   319
lemma (in equivalence) is_closed_eq [dest]:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   320
  "[| x .= x'; x \<in> A; is_closed A; x \<in> carrier S; x' \<in> carrier S |] ==> x' \<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   321
  unfolding eq_is_closed_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   322
  using closure_of_eq [where A = A]
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   323
  by simp
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   324
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   325
lemma (in equivalence) is_closed_eq_rev [dest]:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   326
  "[| x .= x'; x' \<in> A; is_closed A; x \<in> carrier S; x' \<in> carrier S |] ==> x \<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   327
  by (drule sym) (simp_all add: is_closed_eq)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   328
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   329
lemma closure_of_closed [simp, intro]:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   330
  fixes S (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   331
  shows "closure_of A \<subseteq> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   332
unfolding eq_closure_of_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   333
by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   334
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   335
lemma closure_of_memI:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   336
  fixes S (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   337
  assumes "a .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   338
    and "a \<in> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   339
  shows "a \<in> closure_of A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   340
unfolding eq_closure_of_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   341
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   342
by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   343
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   344
lemma closure_ofI2:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   345
  fixes S (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   346
  assumes "a .= a'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   347
    and "a' \<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   348
    and "a \<in> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   349
  shows "a \<in> closure_of A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   350
unfolding eq_closure_of_def elem_def
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   351
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   352
by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   353
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   354
lemma closure_of_memE:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   355
  fixes S (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   356
  assumes p: "a \<in> closure_of A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   357
    and r: "\<lbrakk>a \<in> carrier S; a .\<in> A\<rbrakk> \<Longrightarrow> P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   358
  shows "P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   359
proof -
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   360
  from p
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   361
      have acarr: "a \<in> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   362
      and "a .\<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   363
      by (simp add: eq_closure_of_def)+
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   364
  thus "P" by (rule r)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   365
qed
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   366
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   367
lemma closure_ofE2:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   368
  fixes S (structure)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   369
  assumes p: "a \<in> closure_of A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   370
    and r: "\<And>a'. \<lbrakk>a \<in> carrier S; a' \<in> A; a .= a'\<rbrakk> \<Longrightarrow> P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   371
  shows "P"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   372
proof -
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   373
  from p have acarr: "a \<in> carrier S" by (simp add: eq_closure_of_def)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   374
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   375
  from p have "\<exists>a'\<in>A. a .= a'" by (simp add: eq_closure_of_def elem_def)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   376
  from this obtain a'
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   377
      where "a' \<in> A" and "a .= a'" by auto
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   378
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   379
  from acarr and this
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   380
      show "P" by (rule r)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   381
qed
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   382
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   383
(*
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   384
lemma (in equivalence) classes_consistent:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   385
  assumes Acarr: "A \<subseteq> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   386
  shows "is_closed (closure_of A)"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   387
apply (blast intro: elemI elim elemE)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   388
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   389
apply (intro is_closedI closure_of_memI, simp)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   390
 apply (elim elemE closure_of_memE)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   391
proof -
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   392
  fix x a' a''
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   393
  assume carr: "x \<in> carrier S" "a' \<in> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   394
  assume a''A: "a'' \<in> A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   395
  with Acarr have "a'' \<in> carrier S" by fast
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   396
  note [simp] = carr this Acarr
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   397
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   398
  assume "x .= a'"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   399
  also assume "a' .= a''"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   400
  also from a''A
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   401
       have "a'' .\<in> A" by (simp add: elem_exact)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   402
  finally show "x .\<in> A" by simp
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   403
qed
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   404
*)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   405
(*
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   406
lemma (in equivalence) classes_small:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   407
  assumes "is_closed B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   408
    and "A \<subseteq> B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   409
  shows "closure_of A \<subseteq> B"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   410
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   411
by (blast dest: is_closedD2 elem_subsetD elim: closure_of_memE)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   412
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   413
lemma (in equivalence) classes_eq:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   414
  assumes "A \<subseteq> carrier S"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   415
  shows "A {.=} closure_of A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   416
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   417
by (blast intro: set_eqI elem_exact closure_of_memI elim: closure_of_memE)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   418
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   419
lemma (in equivalence) complete_classes:
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   420
  assumes c: "is_closed A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   421
  shows "A = closure_of A"
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   422
using assms
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   423
by (blast intro: closure_of_memI elem_exact dest: is_closedD1 is_closedD2 closure_of_memE)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   424
*)
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   425
65099
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   426
lemma equivalence_subset:
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   427
  assumes "equivalence L" "A \<subseteq> carrier L"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   428
  shows "equivalence (L\<lparr> carrier := A \<rparr>)"
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   429
proof -
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   430
  interpret L: equivalence L
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   431
    by (simp add: assms)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   432
  show ?thesis
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   433
    by (unfold_locales, simp_all add: L.sym assms rev_subsetD, meson L.trans assms(2) contra_subsetD)
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   434
qed
30d0b2f1df76 Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents: 63167
diff changeset
   435
  
27701
ed7a2e0fab59 New theory on divisibility.
ballarin
parents:
diff changeset
   436
end