src/ZF/equalities.ML
author wenzelm
Tue, 28 Oct 1997 17:36:16 +0100
changeset 4022 0770a19c48d3
parent 2925 b0ae2e13db93
child 4091 771b1f6422a8
permissions -rw-r--r--
added ignored_consts, thms_containing, add_store_axioms(_i), add_store_defs(_i), thms_of; tuned pure thys;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1056
diff changeset
     1
(*  Title:      ZF/equalities
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1056
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Set Theory examples: Union, Intersection, Inclusion, etc.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
    (Thanks also to Philippe de Groote.)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
(** Finite Sets **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
520
806d3f00590d ZF/equalities/Sigma_cons: new
lcp
parents: 517
diff changeset
    12
(* cons_def refers to Upair; reversing the equality LOOPS in rewriting!*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    13
goal ZF.thy "{a} Un B = cons(a,B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    14
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    15
qed "cons_eq";
520
806d3f00590d ZF/equalities/Sigma_cons: new
lcp
parents: 517
diff changeset
    16
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    17
goal ZF.thy "cons(a, cons(b, C)) = cons(b, cons(a, C))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    18
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    19
qed "cons_commute";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    21
goal ZF.thy "!!B. a: B ==> cons(a,B) = B";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    22
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    23
qed "cons_absorb";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    25
goal ZF.thy "!!B. a: B ==> cons(a, B-{a}) = B";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    26
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    27
qed "cons_Diff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    29
goal ZF.thy "!!C. [| a: C;  ALL y:C. y=b |] ==> C = {b}";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    30
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    31
qed "equal_singleton_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
val equal_singleton = ballI RSN (2,equal_singleton_lemma);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
(** Binary Intersection **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
(*NOT an equality, but it seems to belong here...*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    38
goal ZF.thy "cons(a,B) Int C <= cons(a, B Int C)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    39
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    40
qed "Int_cons";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    42
goal ZF.thy "A Int A = A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    43
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    44
qed "Int_absorb";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    46
goal ZF.thy "A Int B = B Int A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    47
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    48
qed "Int_commute";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    49
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    50
goal ZF.thy "(A Int B) Int C  =  A Int (B Int C)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    51
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    52
qed "Int_assoc";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    54
goal ZF.thy "(A Un B) Int C  =  (A Int C) Un (B Int C)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    55
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    56
qed "Int_Un_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    58
goal ZF.thy "A<=B <-> A Int B = A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    59
by (blast_tac (!claset addSEs [equalityE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    60
qed "subset_Int_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    62
goal ZF.thy "A<=B <-> B Int A = A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    63
by (blast_tac (!claset addSEs [equalityE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    64
qed "subset_Int_iff2";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
    65
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    66
goal ZF.thy "!!A B C. C<=A ==> (A-B) Int C = C-B";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    67
by (Blast_tac 1);
1035
279a4fd3c5ce Proved Int_Diff_eq.
lcp
parents: 839
diff changeset
    68
qed "Int_Diff_eq";
279a4fd3c5ce Proved Int_Diff_eq.
lcp
parents: 839
diff changeset
    69
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    70
(** Binary Union **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    72
goal ZF.thy "cons(a,B) Un C = cons(a, B Un C)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    73
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    74
qed "Un_cons";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    76
goal ZF.thy "A Un A = A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    77
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    78
qed "Un_absorb";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    80
goal ZF.thy "A Un B = B Un A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    81
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    82
qed "Un_commute";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    83
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    84
goal ZF.thy "(A Un B) Un C  =  A Un (B Un C)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    85
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    86
qed "Un_assoc";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    88
goal ZF.thy "(A Int B) Un C  =  (A Un C) Int (B Un C)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    89
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    90
qed "Un_Int_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    92
goal ZF.thy "A<=B <-> A Un B = B";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    93
by (blast_tac (!claset addSEs [equalityE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    94
qed "subset_Un_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    96
goal ZF.thy "A<=B <-> B Un A = B";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
    97
by (blast_tac (!claset addSEs [equalityE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
    98
qed "subset_Un_iff2";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
    99
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
(** Simple properties of Diff -- set difference **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   102
goal ZF.thy "A-A = 0";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   103
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   104
qed "Diff_cancel";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   106
goal ZF.thy "0-A = 0";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   107
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   108
qed "empty_Diff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   110
goal ZF.thy "A-0 = A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   111
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   112
qed "Diff_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   113
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   114
goal ZF.thy "A-B=0 <-> A<=B";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   115
by (blast_tac (!claset addEs [equalityE]) 1);
787
1affbb1c5f1f converse_UN, Diff_eq_0_iff: new
lcp
parents: 760
diff changeset
   116
qed "Diff_eq_0_iff";
1affbb1c5f1f converse_UN, Diff_eq_0_iff: new
lcp
parents: 760
diff changeset
   117
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   118
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   119
goal ZF.thy "A - cons(a,B) = A - B - {a}";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   120
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   121
qed "Diff_cons";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   122
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   123
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   124
goal ZF.thy "A - cons(a,B) = A - {a} - B";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   125
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   126
qed "Diff_cons2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   127
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   128
goal ZF.thy "A Int (B-A) = 0";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   129
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   130
qed "Diff_disjoint";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   131
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   132
goal ZF.thy "!!A B. A<=B ==> A Un (B-A) = B";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   133
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   134
qed "Diff_partition";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   135
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   136
goal ZF.thy "A <= B Un (A - B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   137
by (Blast_tac 1);
1611
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   138
qed "subset_Un_Diff";
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   139
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   140
goal ZF.thy "!!A B. [| A<=B; B<=C |] ==> B-(C-A) = A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   141
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   142
qed "double_complement";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   143
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   144
goal ZF.thy "(A Un B) - (B-A) = A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   145
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   146
qed "double_complement_Un";
268
1a038ec6f923 double_complement_Un: new
lcp
parents: 198
diff changeset
   147
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   148
goal ZF.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   149
 "(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)";
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   150
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   151
qed "Un_Int_crazy";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   152
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   153
goal ZF.thy "A - (B Un C) = (A-B) Int (A-C)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   154
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   155
qed "Diff_Un";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   157
goal ZF.thy "A - (B Int C) = (A-B) Un (A-C)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   158
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   159
qed "Diff_Int";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   160
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   161
(*Halmos, Naive Set Theory, page 16.*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   162
goal ZF.thy "(A Int B) Un C = A Int (B Un C)  <->  C<=A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   163
by (blast_tac (!claset addSEs [equalityE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   164
qed "Un_Int_assoc_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   165
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   167
(** Big Union and Intersection **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   168
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   169
goal ZF.thy "Union(cons(a,B)) = a Un Union(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   170
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   171
qed "Union_cons";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   172
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   173
goal ZF.thy "Union(A Un B) = Union(A) Un Union(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   174
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   175
qed "Union_Un_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   177
goal ZF.thy "Union(A Int B) <= Union(A) Int Union(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   178
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   179
qed "Union_Int_subset";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   180
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   181
goal ZF.thy "Union(C) Int A = 0 <-> (ALL B:C. B Int A = 0)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   182
by (blast_tac (!claset addSEs [equalityE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   183
qed "Union_disjoint";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   184
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   185
goalw ZF.thy [Inter_def] "Inter(0) = 0";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   186
by (Blast_tac 1);
1652
9b78ce58d6b1 Proved Inter_0 and converse_INT
paulson
parents: 1611
diff changeset
   187
qed "Inter_0";
9b78ce58d6b1 Proved Inter_0 and converse_INT
paulson
parents: 1611
diff changeset
   188
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   189
goal ZF.thy "!!A B. [| z:A; z:B |] ==> Inter(A) Un Inter(B) <= Inter(A Int B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   190
by (Blast_tac 1);
1568
630d881b51bd New theorem: Inter_Un_subset
paulson
parents: 1461
diff changeset
   191
qed "Inter_Un_subset";
630d881b51bd New theorem: Inter_Un_subset
paulson
parents: 1461
diff changeset
   192
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   193
(* A good challenge: Inter is ill-behaved on the empty set *)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   194
goal ZF.thy "!!A B. [| a:A;  b:B |] ==> Inter(A Un B) = Inter(A) Int Inter(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   195
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   196
qed "Inter_Un_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   197
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   198
goal ZF.thy "Union({b}) = b";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   199
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   200
qed "Union_singleton";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   201
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   202
goal ZF.thy "Inter({b}) = b";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   203
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   204
qed "Inter_singleton";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   205
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   206
(** Unions and Intersections of Families **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   207
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   208
goal ZF.thy "Union(A) = (UN x:A. x)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   209
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   210
qed "Union_eq_UN";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   211
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   212
goalw ZF.thy [Inter_def] "Inter(A) = (INT x:A. x)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   213
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   214
qed "Inter_eq_INT";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   215
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   216
goal ZF.thy "(UN i:0. A(i)) = 0";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   217
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   218
qed "UN_0";
517
a9f93400f307 for infinite datatypes with arbitrary index sets
lcp
parents: 435
diff changeset
   219
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   220
(*Halmos, Naive Set Theory, page 35.*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   221
goal ZF.thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   222
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   223
qed "Int_UN_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   224
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   225
goal ZF.thy "!!A B. i:I ==> B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   226
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   227
qed "Un_INT_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   228
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   229
goal ZF.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   230
    "(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))";
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   231
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   232
qed "Int_UN_distrib2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   233
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   234
goal ZF.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   235
    "!!I J. [| i:I;  j:J |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   236
\    (INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))";
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   237
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   238
qed "Un_INT_distrib2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   239
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   240
goal ZF.thy "!!A. a: A ==> (UN y:A. c) = c";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   241
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   242
qed "UN_constant";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   243
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   244
goal ZF.thy "!!A. a: A ==> (INT y:A. c) = c";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   245
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   246
qed "INT_constant";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   247
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   248
(** Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   249
    Union of a family of unions **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   250
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   251
goal ZF.thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i))  Un  (UN i:I. B(i))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   252
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   253
qed "UN_Un_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   254
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   255
goal ZF.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   256
    "!!A B. i:I ==> \
192
3dc5c8016a0e ZF/equalities/SUM_eq_UN: new
lcp
parents: 182
diff changeset
   257
\           (INT i:I. A(i)  Int  B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))";
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   258
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   259
qed "INT_Int_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   260
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   261
goal ZF.thy "(UN z:I Int J. A(z)) <= (UN z:I. A(z)) Int (UN z:J. A(z))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   262
by (Blast_tac 1);
1784
036a7f301623 Added a new theorem, UN_Int_subset
paulson
parents: 1652
diff changeset
   263
qed "UN_Int_subset";
036a7f301623 Added a new theorem, UN_Int_subset
paulson
parents: 1652
diff changeset
   264
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   265
(** Devlin, page 12, exercise 5: Complements **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   266
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   267
goal ZF.thy "!!A B. i:I ==> B - (UN i:I. A(i)) = (INT i:I. B - A(i))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   268
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   269
qed "Diff_UN";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   270
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   271
goal ZF.thy "!!A B. i:I ==> B - (INT i:I. A(i)) = (UN i:I. B - A(i))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   272
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   273
qed "Diff_INT";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   274
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   275
(** Unions and Intersections with General Sum **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   276
1611
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   277
(*Not suitable for rewriting: LOOPS!*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   278
goal ZF.thy "Sigma(cons(a,B), C) = ({a}*C(a)) Un Sigma(B,C)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   279
by (Blast_tac 1);
1611
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   280
qed "Sigma_cons1";
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   281
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   282
(*Not suitable for rewriting: LOOPS!*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   283
goal ZF.thy "A * cons(b,B) = A*{b} Un A*B";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   284
by (Blast_tac 1);
1611
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   285
qed "Sigma_cons2";
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   286
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   287
goal ZF.thy "Sigma(succ(A), B) = ({A}*B(A)) Un Sigma(A,B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   288
by (Blast_tac 1);
1611
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   289
qed "Sigma_succ1";
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   290
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   291
goal ZF.thy "A * succ(B) = A*{B} Un A*B";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   292
by (Blast_tac 1);
1611
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   293
qed "Sigma_succ2";
520
806d3f00590d ZF/equalities/Sigma_cons: new
lcp
parents: 517
diff changeset
   294
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   295
goal ZF.thy "(SUM x:(UN y:A. C(y)). B(x)) = (UN y:A. SUM x:C(y). B(x))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   296
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   297
qed "SUM_UN_distrib1";
182
e30b55c07235 New distributive laws for Sigma and UN
lcp
parents: 0
diff changeset
   298
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   299
goal ZF.thy "(SUM i:I. UN j:J. C(i,j)) = (UN j:J. SUM i:I. C(i,j))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   300
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   301
qed "SUM_UN_distrib2";
182
e30b55c07235 New distributive laws for Sigma and UN
lcp
parents: 0
diff changeset
   302
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   303
goal ZF.thy "(SUM i:I Un J. C(i)) = (SUM i:I. C(i)) Un (SUM j:J. C(j))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   304
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   305
qed "SUM_Un_distrib1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   306
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   307
goal ZF.thy "(SUM i:I. A(i) Un B(i)) = (SUM i:I. A(i)) Un (SUM i:I. B(i))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   308
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   309
qed "SUM_Un_distrib2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   310
685
0727f0c0c4f0 ZF/equalities/domain_converse,range_converse,
lcp
parents: 536
diff changeset
   311
(*First-order version of the above, for rewriting*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   312
goal ZF.thy "I * (A Un B) = I*A Un I*B";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1056
diff changeset
   313
by (rtac SUM_Un_distrib2 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   314
qed "prod_Un_distrib2";
685
0727f0c0c4f0 ZF/equalities/domain_converse,range_converse,
lcp
parents: 536
diff changeset
   315
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   316
goal ZF.thy "(SUM i:I Int J. C(i)) = (SUM i:I. C(i)) Int (SUM j:J. C(j))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   317
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   318
qed "SUM_Int_distrib1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   319
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   320
goal ZF.thy
192
3dc5c8016a0e ZF/equalities/SUM_eq_UN: new
lcp
parents: 182
diff changeset
   321
    "(SUM i:I. A(i) Int B(i)) = (SUM i:I. A(i)) Int (SUM i:I. B(i))";
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   322
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   323
qed "SUM_Int_distrib2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   324
685
0727f0c0c4f0 ZF/equalities/domain_converse,range_converse,
lcp
parents: 536
diff changeset
   325
(*First-order version of the above, for rewriting*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   326
goal ZF.thy "I * (A Int B) = I*A Int I*B";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1056
diff changeset
   327
by (rtac SUM_Int_distrib2 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   328
qed "prod_Int_distrib2";
685
0727f0c0c4f0 ZF/equalities/domain_converse,range_converse,
lcp
parents: 536
diff changeset
   329
192
3dc5c8016a0e ZF/equalities/SUM_eq_UN: new
lcp
parents: 182
diff changeset
   330
(*Cf Aczel, Non-Well-Founded Sets, page 115*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   331
goal ZF.thy "(SUM i:I. A(i)) = (UN i:I. {i} * A(i))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   332
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   333
qed "SUM_eq_UN";
192
3dc5c8016a0e ZF/equalities/SUM_eq_UN: new
lcp
parents: 182
diff changeset
   334
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   335
(** Domain **)
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   336
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   337
qed_goal "domain_of_prod" ZF.thy "!!A B. b:B ==> domain(A*B) = A"
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   338
 (fn _ => [ Blast_tac 1 ]);
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   339
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   340
qed_goal "domain_0" ZF.thy "domain(0) = 0"
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   341
 (fn _ => [ Blast_tac 1 ]);
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   342
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   343
qed_goal "domain_cons" ZF.thy
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   344
    "domain(cons(<a,b>,r)) = cons(a, domain(r))"
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   345
 (fn _ => [ Blast_tac 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   346
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   347
goal ZF.thy "domain(A Un B) = domain(A) Un domain(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   348
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   349
qed "domain_Un_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   350
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   351
goal ZF.thy "domain(A Int B) <= domain(A) Int domain(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   352
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   353
qed "domain_Int_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   354
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   355
goal ZF.thy "domain(A) - domain(B) <= domain(A - B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   356
by (Blast_tac 1);
1056
097b3255bf3a Renamed domain_diff_subset, range_diff_subset,
lcp
parents: 1035
diff changeset
   357
qed "domain_Diff_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   358
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   359
goal ZF.thy "domain(converse(r)) = range(r)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   360
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   361
qed "domain_converse";
685
0727f0c0c4f0 ZF/equalities/domain_converse,range_converse,
lcp
parents: 536
diff changeset
   362
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   363
Addsimps [domain_0, domain_cons, domain_Un_eq, domain_converse];
685
0727f0c0c4f0 ZF/equalities/domain_converse,range_converse,
lcp
parents: 536
diff changeset
   364
0727f0c0c4f0 ZF/equalities/domain_converse,range_converse,
lcp
parents: 536
diff changeset
   365
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   366
(** Range **)
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   367
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   368
qed_goal "range_of_prod" ZF.thy
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   369
    "!!a A B. a:A ==> range(A*B) = B"
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   370
 (fn _ => [ Blast_tac 1 ]);
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   371
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   372
qed_goal "range_0" ZF.thy "range(0) = 0"
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   373
 (fn _ => [ Blast_tac 1 ]); 
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   374
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   375
qed_goal "range_cons" ZF.thy
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   376
    "range(cons(<a,b>,r)) = cons(b, range(r))"
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   377
 (fn _ => [ Blast_tac 1 ]);
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   378
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   379
goal ZF.thy "range(A Un B) = range(A) Un range(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   380
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   381
qed "range_Un_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   382
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   383
goal ZF.thy "range(A Int B) <= range(A) Int range(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   384
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   385
qed "range_Int_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   386
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   387
goal ZF.thy "range(A) - range(B) <= range(A - B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   388
by (Blast_tac 1);
1056
097b3255bf3a Renamed domain_diff_subset, range_diff_subset,
lcp
parents: 1035
diff changeset
   389
qed "range_Diff_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   390
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   391
goal ZF.thy "range(converse(r)) = domain(r)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   392
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   393
qed "range_converse";
685
0727f0c0c4f0 ZF/equalities/domain_converse,range_converse,
lcp
parents: 536
diff changeset
   394
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   395
Addsimps [range_0, range_cons, range_Un_eq, range_converse];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   396
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   397
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   398
(** Field **)
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   399
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   400
qed_goal "field_of_prod" ZF.thy "field(A*A) = A"
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   401
 (fn _ => [ Blast_tac 1 ]); 
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   402
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   403
qed_goal "field_0" ZF.thy "field(0) = 0"
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   404
 (fn _ => [ Blast_tac 1 ]); 
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   405
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   406
qed_goal "field_cons" ZF.thy
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   407
    "field(cons(<a,b>,r)) = cons(a, cons(b, field(r)))"
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   408
 (fn _ => [ rtac equalityI 1, ALLGOALS (Blast_tac) ]);
536
5fbfa997f1b0 ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents: 520
diff changeset
   409
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   410
goal ZF.thy "field(A Un B) = field(A) Un field(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   411
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   412
qed "field_Un_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   413
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   414
goal ZF.thy "field(A Int B) <= field(A) Int field(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   415
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   416
qed "field_Int_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   417
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   418
goal ZF.thy "field(A) - field(B) <= field(A - B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   419
by (Blast_tac 1);
1056
097b3255bf3a Renamed domain_diff_subset, range_diff_subset,
lcp
parents: 1035
diff changeset
   420
qed "field_Diff_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   421
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   422
goal ZF.thy "field(converse(r)) = field(r)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   423
by (Blast_tac 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   424
qed "field_converse";
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   425
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   426
Addsimps [field_0, field_cons, field_Un_eq, field_converse];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   427
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   428
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   429
(** Image **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   430
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   431
goal ZF.thy "r``0 = 0";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   432
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   433
qed "image_0";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   434
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   435
goal ZF.thy "r``(A Un B) = (r``A) Un (r``B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   436
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   437
qed "image_Un";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   438
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   439
goal ZF.thy "r``(A Int B) <= (r``A) Int (r``B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   440
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   441
qed "image_Int_subset";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   442
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   443
goal ZF.thy "(r Int A*A)``B <= (r``B) Int A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   444
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   445
qed "image_Int_square_subset";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   446
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   447
goal ZF.thy "!!r. B<=A ==> (r Int A*A)``B = (r``B) Int A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   448
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   449
qed "image_Int_square";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   450
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   451
Addsimps [image_0, image_Un];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   452
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   453
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   454
(** Inverse Image **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   455
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   456
goal ZF.thy "r-``0 = 0";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   457
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   458
qed "vimage_0";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   459
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   460
goal ZF.thy "r-``(A Un B) = (r-``A) Un (r-``B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   461
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   462
qed "vimage_Un";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   463
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   464
goal ZF.thy "r-``(A Int B) <= (r-``A) Int (r-``B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   465
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   466
qed "vimage_Int_subset";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   467
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   468
goal ZF.thy "(r Int A*A)-``B <= (r-``B) Int A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   469
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   470
qed "vimage_Int_square_subset";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   471
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   472
goal ZF.thy "!!r. B<=A ==> (r Int A*A)-``B = (r-``B) Int A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   473
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   474
qed "vimage_Int_square";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   475
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   476
Addsimps [vimage_0, vimage_Un];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   477
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   478
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   479
(** Converse **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   480
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   481
goal ZF.thy "converse(A Un B) = converse(A) Un converse(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   482
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   483
qed "converse_Un";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   484
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   485
goal ZF.thy "converse(A Int B) = converse(A) Int converse(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   486
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   487
qed "converse_Int";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   488
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   489
goal ZF.thy "converse(A - B) = converse(A) - converse(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   490
by (Blast_tac 1);
1056
097b3255bf3a Renamed domain_diff_subset, range_diff_subset,
lcp
parents: 1035
diff changeset
   491
qed "converse_Diff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   492
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   493
goal ZF.thy "converse(UN x:A. B(x)) = (UN x:A. converse(B(x)))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   494
by (Blast_tac 1);
787
1affbb1c5f1f converse_UN, Diff_eq_0_iff: new
lcp
parents: 760
diff changeset
   495
qed "converse_UN";
1affbb1c5f1f converse_UN, Diff_eq_0_iff: new
lcp
parents: 760
diff changeset
   496
1652
9b78ce58d6b1 Proved Inter_0 and converse_INT
paulson
parents: 1611
diff changeset
   497
(*Unfolding Inter avoids using excluded middle on A=0*)
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   498
goalw ZF.thy [Inter_def] "converse(INT x:A. B(x)) = (INT x:A. converse(B(x)))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   499
by (Blast_tac 1);
1652
9b78ce58d6b1 Proved Inter_0 and converse_INT
paulson
parents: 1611
diff changeset
   500
qed "converse_INT";
9b78ce58d6b1 Proved Inter_0 and converse_INT
paulson
parents: 1611
diff changeset
   501
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   502
Addsimps [converse_Un, converse_Int, converse_Diff, converse_UN, converse_INT];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   503
198
0f0ff91b07f6 new section for equality properties
lcp
parents: 192
diff changeset
   504
(** Pow **)
0f0ff91b07f6 new section for equality properties
lcp
parents: 192
diff changeset
   505
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   506
goal ZF.thy "Pow(A) Un Pow(B) <= Pow(A Un B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   507
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   508
qed "Un_Pow_subset";
198
0f0ff91b07f6 new section for equality properties
lcp
parents: 192
diff changeset
   509
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   510
goal ZF.thy "(UN x:A. Pow(B(x))) <= Pow(UN x:A. B(x))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   511
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   512
qed "UN_Pow_subset";
198
0f0ff91b07f6 new section for equality properties
lcp
parents: 192
diff changeset
   513
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   514
goal ZF.thy "A <= Pow(Union(A))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   515
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   516
qed "subset_Pow_Union";
198
0f0ff91b07f6 new section for equality properties
lcp
parents: 192
diff changeset
   517
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   518
goal ZF.thy "Union(Pow(A)) = A";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   519
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   520
qed "Union_Pow_eq";
198
0f0ff91b07f6 new section for equality properties
lcp
parents: 192
diff changeset
   521
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   522
goal ZF.thy "Pow(A Int B) = Pow(A) Int Pow(B)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   523
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   524
qed "Int_Pow_eq";
198
0f0ff91b07f6 new section for equality properties
lcp
parents: 192
diff changeset
   525
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   526
goal ZF.thy "!!x A. x:A ==> (INT x:A. Pow(B(x))) = Pow(INT x:A. B(x))";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   527
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 685
diff changeset
   528
qed "INT_Pow_subset";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 268
diff changeset
   529
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   530
Addsimps [Union_Pow_eq, Int_Pow_eq];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   531
839
1aa6b351ca34 RepFun_eq_0_iff, RepFun_0: new
lcp
parents: 787
diff changeset
   532
(** RepFun **)
1aa6b351ca34 RepFun_eq_0_iff, RepFun_0: new
lcp
parents: 787
diff changeset
   533
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   534
goal ZF.thy "{f(x).x:A}=0 <-> A=0";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2877
diff changeset
   535
	(*blast_tac takes too long to find a good depth*)
b0ae2e13db93 Using Blast_tac
paulson
parents: 2877
diff changeset
   536
by (Blast.depth_tac (!claset addSEs [equalityE]) 10 1);
839
1aa6b351ca34 RepFun_eq_0_iff, RepFun_0: new
lcp
parents: 787
diff changeset
   537
qed "RepFun_eq_0_iff";
1aa6b351ca34 RepFun_eq_0_iff, RepFun_0: new
lcp
parents: 787
diff changeset
   538
1611
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   539
(** Collect **)
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   540
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   541
goal ZF.thy "Collect(A Un B, P) = Collect(A,P) Un Collect(B,P)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   542
by (Blast_tac 1);
1611
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   543
qed "Collect_Un";
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   544
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   545
goal ZF.thy "Collect(A Int B, P) = Collect(A,P) Int Collect(B,P)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   546
by (Blast_tac 1);
1611
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   547
qed "Collect_Int";
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   548
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   549
goal ZF.thy "Collect(A - B, P) = Collect(A,P) - Collect(B,P)";
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   550
by (Blast_tac 1);
1611
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   551
qed "Collect_Diff";
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   552
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   553
goal ZF.thy "{x:cons(a,B). P(x)} = if(P(a), cons(a, {x:B. P(x)}), {x:B. P(x)})";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1784
diff changeset
   554
by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2493
diff changeset
   555
by (Blast_tac 1);
1611
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   556
qed "Collect_cons";
35e0fd1b1775 New results from AC/Cardinal_aux.ML
paulson
parents: 1568
diff changeset
   557