src/ZF/EquivClass.thy
author wenzelm
Wed, 10 Feb 2016 14:14:43 +0100
changeset 62278 c04e97be39d3
parent 61798 27f3c10b0b50
child 67443 3abf6a722518
permissions -rw-r--r--
misc tuning and updates;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     1
(*  Title:      ZF/EquivClass.thy
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     3
    Copyright   1994  University of Cambridge
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     4
*)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     5
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
     6
section\<open>Equivalence Relations\<close>
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     7
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     8
theory EquivClass imports Trancl Perm begin
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     9
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    10
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    11
  quotient   :: "[i,i]=>i"    (infixl "'/'/" 90)  (*set of equiv classes*)  where
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    12
      "A//r == {r``{x} . x \<in> A}"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    13
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    14
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    15
  congruent  :: "[i,i=>i]=>o"  where
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    16
      "congruent(r,b) == \<forall>y z. <y,z>:r \<longrightarrow> b(y)=b(z)"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    17
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    18
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    19
  congruent2 :: "[i,i,[i,i]=>i]=>o"  where
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    20
      "congruent2(r1,r2,b) == \<forall>y1 z1 y2 z2.
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    21
           <y1,z1>:r1 \<longrightarrow> <y2,z2>:r2 \<longrightarrow> b(y1,y2) = b(z1,z2)"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    22
24892
c663e675e177 replaced some 'translations' by 'abbreviation';
wenzelm
parents: 23146
diff changeset
    23
abbreviation
c663e675e177 replaced some 'translations' by 'abbreviation';
wenzelm
parents: 23146
diff changeset
    24
  RESPECTS ::"[i=>i, i] => o"  (infixr "respects" 80) where
c663e675e177 replaced some 'translations' by 'abbreviation';
wenzelm
parents: 23146
diff changeset
    25
  "f respects r == congruent(r,f)"
c663e675e177 replaced some 'translations' by 'abbreviation';
wenzelm
parents: 23146
diff changeset
    26
c663e675e177 replaced some 'translations' by 'abbreviation';
wenzelm
parents: 23146
diff changeset
    27
abbreviation
c663e675e177 replaced some 'translations' by 'abbreviation';
wenzelm
parents: 23146
diff changeset
    28
  RESPECTS2 ::"[i=>i=>i, i] => o"  (infixr "respects2 " 80) where
c663e675e177 replaced some 'translations' by 'abbreviation';
wenzelm
parents: 23146
diff changeset
    29
  "f respects2 r == congruent2(r,r,f)"
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
    30
    \<comment>\<open>Abbreviation for the common case where the relations are identical\<close>
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    31
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    32
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
    33
subsection\<open>Suppes, Theorem 70:
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
    34
    @{term r} is an equiv relation iff @{term "converse(r) O r = r"}\<close>
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    35
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    36
(** first half: equiv(A,r) ==> converse(r) O r = r **)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    37
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    38
lemma sym_trans_comp_subset:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    39
    "[| sym(r); trans(r) |] ==> converse(r) O r \<subseteq> r"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    40
by (unfold trans_def sym_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    41
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    42
lemma refl_comp_subset:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    43
    "[| refl(A,r); r \<subseteq> A*A |] ==> r \<subseteq> converse(r) O r"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    44
by (unfold refl_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    45
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    46
lemma equiv_comp_eq:
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    47
    "equiv(A,r) ==> converse(r) O r = r"
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    48
apply (unfold equiv_def)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    49
apply (blast del: subsetI intro!: sym_trans_comp_subset refl_comp_subset)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    50
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    51
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    52
(*second half*)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    53
lemma comp_equivI:
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    54
    "[| converse(r) O r = r;  domain(r) = A |] ==> equiv(A,r)"
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    55
apply (unfold equiv_def refl_def sym_def trans_def)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    56
apply (erule equalityE)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    57
apply (subgoal_tac "\<forall>x y. <x,y> \<in> r \<longrightarrow> <y,x> \<in> r", blast+)
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    58
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    59
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    60
(** Equivalence classes **)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    61
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    62
(*Lemma for the next result*)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    63
lemma equiv_class_subset:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    64
    "[| sym(r);  trans(r);  <a,b>: r |] ==> r``{a} \<subseteq> r``{b}"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    65
by (unfold trans_def sym_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    66
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    67
lemma equiv_class_eq:
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    68
    "[| equiv(A,r);  <a,b>: r |] ==> r``{a} = r``{b}"
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    69
apply (unfold equiv_def)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    70
apply (safe del: subsetI intro!: equalityI equiv_class_subset)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    71
apply (unfold sym_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    72
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    73
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    74
lemma equiv_class_self:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    75
    "[| equiv(A,r);  a \<in> A |] ==> a \<in> r``{a}"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    76
by (unfold equiv_def refl_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    77
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    78
(*Lemma for the next result*)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    79
lemma subset_equiv_class:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    80
    "[| equiv(A,r);  r``{b} \<subseteq> r``{a};  b \<in> A |] ==> <a,b>: r"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    81
by (unfold equiv_def refl_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    82
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    83
lemma eq_equiv_class: "[| r``{a} = r``{b};  equiv(A,r);  b \<in> A |] ==> <a,b>: r"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    84
by (assumption | rule equalityD2 subset_equiv_class)+
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    85
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    86
(*thus r``{a} = r``{b} as well*)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    87
lemma equiv_class_nondisjoint:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    88
    "[| equiv(A,r);  x: (r``{a} \<inter> r``{b}) |] ==> <a,b>: r"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    89
by (unfold equiv_def trans_def sym_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    90
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
    91
lemma equiv_type: "equiv(A,r) ==> r \<subseteq> A*A"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    92
by (unfold equiv_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    93
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    94
lemma equiv_class_eq_iff:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    95
     "equiv(A,r) ==> <x,y>: r \<longleftrightarrow> r``{x} = r``{y} & x \<in> A & y \<in> A"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    96
by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    97
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    98
lemma eq_equiv_class_iff:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
    99
     "[| equiv(A,r);  x \<in> A;  y \<in> A |] ==> r``{x} = r``{y} \<longleftrightarrow> <x,y>: r"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   100
by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   101
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   102
(*** Quotients ***)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   103
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   104
(** Introduction/elimination rules -- needed? **)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   105
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   106
lemma quotientI [TC]: "x \<in> A ==> r``{x}: A//r"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   107
apply (unfold quotient_def)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   108
apply (erule RepFunI)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   109
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   110
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   111
lemma quotientE:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   112
    "[| X \<in> A//r;  !!x. [| X = r``{x};  x \<in> A |] ==> P |] ==> P"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   113
by (unfold quotient_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   114
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   115
lemma Union_quotient:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   116
    "equiv(A,r) ==> \<Union>(A//r) = A"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   117
by (unfold equiv_def refl_def quotient_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   118
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   119
lemma quotient_disj:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   120
    "[| equiv(A,r);  X \<in> A//r;  Y \<in> A//r |] ==> X=Y | (X \<inter> Y \<subseteq> 0)"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   121
apply (unfold quotient_def)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   122
apply (safe intro!: equiv_class_eq, assumption)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   123
apply (unfold equiv_def trans_def sym_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   124
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   125
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   126
subsection\<open>Defining Unary Operations upon Equivalence Classes\<close>
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   127
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   128
(** Could have a locale with the premises equiv(A,r)  and  congruent(r,b)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   129
**)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   130
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   131
(*Conversion rule*)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   132
lemma UN_equiv_class:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   133
    "[| equiv(A,r);  b respects r;  a \<in> A |] ==> (\<Union>x\<in>r``{a}. b(x)) = b(a)"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   134
apply (subgoal_tac "\<forall>x \<in> r``{a}. b(x) = b(a)")
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   135
 apply simp
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   136
 apply (blast intro: equiv_class_self)
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   137
apply (unfold equiv_def sym_def congruent_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   138
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   139
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   140
(*type checking of  @{term"\<Union>x\<in>r``{a}. b(x)"} *)
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   141
lemma UN_equiv_class_type:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   142
    "[| equiv(A,r);  b respects r;  X \<in> A//r;  !!x.  x \<in> A ==> b(x) \<in> B |]
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   143
     ==> (\<Union>x\<in>X. b(x)) \<in> B"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   144
apply (unfold quotient_def, safe)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   145
apply (simp (no_asm_simp) add: UN_equiv_class)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   146
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   147
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   148
(*Sufficient conditions for injectiveness.  Could weaken premises!
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   149
  major premise could be an inclusion; bcong could be !!y. y \<in> A ==> b(y):B
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   150
*)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   151
lemma UN_equiv_class_inject:
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   152
    "[| equiv(A,r);   b respects r;
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   153
        (\<Union>x\<in>X. b(x))=(\<Union>y\<in>Y. b(y));  X \<in> A//r;  Y \<in> A//r;
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   154
        !!x y. [| x \<in> A; y \<in> A; b(x)=b(y) |] ==> <x,y>:r |]
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   155
     ==> X=Y"
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   156
apply (unfold quotient_def, safe)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   157
apply (rule equiv_class_eq, assumption)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   158
apply (simp add: UN_equiv_class [of A r b])
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   159
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   160
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   161
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   162
subsection\<open>Defining Binary Operations upon Equivalence Classes\<close>
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   163
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   164
lemma congruent2_implies_congruent:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   165
    "[| equiv(A,r1);  congruent2(r1,r2,b);  a \<in> A |] ==> congruent(r2,b(a))"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   166
by (unfold congruent_def congruent2_def equiv_def refl_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   167
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   168
lemma congruent2_implies_congruent_UN:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   169
    "[| equiv(A1,r1);  equiv(A2,r2);  congruent2(r1,r2,b);  a \<in> A2 |] ==>
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   170
     congruent(r1, %x1. \<Union>x2 \<in> r2``{a}. b(x1,x2))"
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   171
apply (unfold congruent_def, safe)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   172
apply (frule equiv_type [THEN subsetD], assumption)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   173
apply clarify
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   174
apply (simp add: UN_equiv_class congruent2_implies_congruent)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   175
apply (unfold congruent2_def equiv_def refl_def, blast)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   176
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   177
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   178
lemma UN_equiv_class2:
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   179
    "[| equiv(A1,r1);  equiv(A2,r2);  congruent2(r1,r2,b);  a1: A1;  a2: A2 |]
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   180
     ==> (\<Union>x1 \<in> r1``{a1}. \<Union>x2 \<in> r2``{a2}. b(x1,x2)) = b(a1,a2)"
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   181
by (simp add: UN_equiv_class congruent2_implies_congruent
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   182
              congruent2_implies_congruent_UN)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   183
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   184
(*type checking*)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   185
lemma UN_equiv_class_type2:
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   186
    "[| equiv(A,r);  b respects2 r;
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   187
        X1: A//r;  X2: A//r;
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   188
        !!x1 x2.  [| x1: A; x2: A |] ==> b(x1,x2) \<in> B
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   189
     |] ==> (\<Union>x1\<in>X1. \<Union>x2\<in>X2. b(x1,x2)) \<in> B"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   190
apply (unfold quotient_def, safe)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   191
apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   192
                    congruent2_implies_congruent quotientI)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   193
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   194
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   195
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   196
(*Suggested by John Harrison -- the two subproofs may be MUCH simpler
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   197
  than the direct proof*)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   198
lemma congruent2I:
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   199
    "[|  equiv(A1,r1);  equiv(A2,r2);
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   200
        !! y z w. [| w \<in> A2;  <y,z> \<in> r1 |] ==> b(y,w) = b(z,w);
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   201
        !! y z w. [| w \<in> A1;  <y,z> \<in> r2 |] ==> b(w,y) = b(w,z)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   202
     |] ==> congruent2(r1,r2,b)"
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   203
apply (unfold congruent2_def equiv_def refl_def, safe)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   204
apply (blast intro: trans)
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   205
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   206
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   207
lemma congruent2_commuteI:
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   208
 assumes equivA: "equiv(A,r)"
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   209
     and commute: "!! y z. [| y \<in> A;  z \<in> A |] ==> b(y,z) = b(z,y)"
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   210
     and congt:   "!! y z w. [| w \<in> A;  <y,z>: r |] ==> b(w,y) = b(w,z)"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   211
 shows "b respects2 r"
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   212
apply (insert equivA [THEN equiv_type, THEN subsetD])
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   213
apply (rule congruent2I [OF equivA equivA])
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   214
apply (rule commute [THEN trans])
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   215
apply (rule_tac [3] commute [THEN trans, symmetric])
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   216
apply (rule_tac [5] sym)
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   217
apply (blast intro: congt)+
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   218
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   219
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   220
(*Obsolete?*)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   221
lemma congruent_commuteI:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   222
    "[| equiv(A,r);  Z \<in> A//r;
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   223
        !!w. [| w \<in> A |] ==> congruent(r, %z. b(w,z));
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46821
diff changeset
   224
        !!x y. [| x \<in> A;  y \<in> A |] ==> b(y,x) = b(x,y)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   225
     |] ==> congruent(r, %w. \<Union>z\<in>Z. b(w,z))"
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   226
apply (simp (no_asm) add: congruent_def)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   227
apply (safe elim!: quotientE)
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   228
apply (frule equiv_type [THEN subsetD], assumption)
46820
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   229
apply (simp add: UN_equiv_class [of A r])
c656222c4dc1 mathematical symbols instead of ASCII
paulson
parents: 35762
diff changeset
   230
apply (simp add: congruent_def)
23146
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   231
done
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   232
0bc590051d95 moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   233
end