src/ZF/UNITY/UNITYMisc.ML
author wenzelm
Wed, 31 Oct 2001 01:21:01 +0100
changeset 11990 c1daefc08eff
parent 11479 697dcaaf478f
child 12215 55c084921240
permissions -rw-r--r--
use induct_rulify2;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/UNITYMisc.ML
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Sidi O Ehmety, Computer Laboratory
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   2001  University of Cambridge
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     5
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     6
Some miscellaneous and add-hoc set theory concepts.
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     7
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     8
*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     9
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    10
Goalw [measure_def, less_than_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    11
         "less_than(A) = {<x,y>:A*A. x<y}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    12
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    13
qed "less_than_equals";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    14
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    15
Goalw [less_than_def] "wf(less_than(A))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    16
by (rtac wf_measure 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    17
qed "wf_less_than";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    18
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    19
Goalw [less_than_def, measure_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    20
       "less_than(A)<= A*A";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    21
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    22
qed "less_than_subset";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    23
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    24
Goalw [less_than_def, measure_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    25
"<x,y>:less_than(A) <-> (x:A & y:A & x<y)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    26
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    27
qed "less_than_iff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    28
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    29
Goalw [lessThan_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    30
"i:lessThan(k,A) <-> i:A & i<k";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    31
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    32
qed "lessThan_iff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    33
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    34
Goalw [greaterThan_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    35
"i:greaterThan(k,A) <-> i:A & k<i";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    36
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    37
qed "greaterThan_iff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    38
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    39
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    40
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    41
(** Needed for WF reasoning in WFair.ML **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    42
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    43
Goal "k:A ==>less_than(A)``{k} = greaterThan(k, A)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    44
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    45
by (auto_tac (claset(), simpset() addsimps 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    46
                [less_than_iff,greaterThan_iff]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    47
qed "Image_less_than";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    48
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    49
Goal "k:A ==> less_than(A)-`` {k} = lessThan(k, A)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    50
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    51
by (auto_tac (claset(), simpset() addsimps 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    52
                [less_than_iff,lessThan_iff]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    53
qed "Image_inverse_less_than";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    54
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    55
Addsimps [Image_less_than, Image_inverse_less_than];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    56
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    57
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    58
(** Ad-hoc set-theory rules **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    59
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    60
Goal "(C Int A) Un (C Int B) = C Int (A Un B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    61
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    62
qed "Int_Un_distrib2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    63
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    64
Goal "C Int (A - B) = (C Int A) - (C Int B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    65
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    66
qed "Diff_Int_distrib";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    67
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    68
Goal "(A - B) Int C = (A Int C) - (B Int C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    69
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    70
qed "Diff_Int_distrib2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    71
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    72
Goal "A Un (A Un B) = A Un B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    73
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    74
qed "double_Un";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    75
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    76
Goal "A Un (B Un C) = B Un (A Un C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    77
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    78
qed "Un_assoc2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    79
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    80
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    81
val Un_ac = [Un_assoc, double_Un, Un_commute, Un_assoc2, Un_absorb];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    82
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    83
Goal "A Un B = Union({A, B})";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    84
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    85
qed "Un_eq_Union";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    86
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    87
Goal "(UN x:A. {x}) = A";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    88
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    89
qed "UN_singleton";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    90
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    91
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    92
Goal "Union(B) Int A = (UN b:B. b Int A)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    93
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    94
qed "Int_Union_Union";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    95
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    96
Goal "A Int Union(B) = (UN b:B. A Int b)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    97
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    98
qed "Int_Union_Union2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    99
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   100
(** Intersection **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   101
Goal "A Int (A Int B) = A Int B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   102
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   103
qed "double_Int";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   104
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   105
Goal "A Int (B Int C) = B Int (A Int C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   106
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   107
qed "Int_assoc2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   108
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   109
val Int_ac = [Int_assoc, double_Int, Int_commute, Int_assoc2];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   110
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   111
Goal "A  Int B = 0 ==> A - B = A";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   112
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   113
qed "Diff_triv";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   114
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   115
Goal "A Int B - C = A Int (B - C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   116
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   117
qed "Int_Diff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   118
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   119
Goal "f -``B = (UN y:B. f-``{y})";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   120
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   121
qed "vimage_eq_UN";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   122
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   123
Goal "A Un B - (A - B) = B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   124
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   125
qed "Un_Diff_Diff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   126
AddIffs [Un_Diff_Diff];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   127
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   128
(*** INT simps ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   129
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   130
Goal 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   131
"i:I ==> (INT i:I. A(i)) Un B = (INT i:I. A(i) Un B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   132
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   133
qed "INT_Un_distrib";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   134
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   135
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   136
 "!!C. c: C ==> (INT x:C. cons(a, B(x))) = cons(a, (INT x:C. B(x)))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   137
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   138
qed "INT_cons";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   139
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   140
Goal 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   141
"i:I ==> (INT i:I. A(i)) Int B = (INT i:I. A(i) Int B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   142
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   143
qed "INT_Int_distrib2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   144
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   145
Goal 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   146
"i:I ==> B Int (INT i:I. A(i)) = (INT i:I. B Int A(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   147
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   148
qed "Int_INT_distrib";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   149
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   150
val INT_simps = [Un_INT_distrib, Un_INT_distrib2, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   151
                 INT_constant, INT_Int_distrib, Diff_INT, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   152
                 INT_Un_distrib,INT_cons, INT_Int_distrib2, Int_INT_distrib, Inter_0];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   153
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   154
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   155
(*** UN ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   156
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   157
"!!C. c: C ==> cons(a, (UN x:C. B(x))) = (UN x:C. cons(a, B(x)))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   158
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   159
qed "UN_cons";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   160
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   161
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   162
 "!!C. c: C ==> B Un (UN x:C. A(x)) = (UN x:C. B  Un A(x))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   163
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   164
qed "Un_UN_distrib";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   165
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   166
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   167
"!!C. c: C ==> (UN x:C. B(x)) Un A   =  (UN x:C. B(x) Un A)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   168
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   169
qed "UN_Un_distrib2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   170
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   171
Goal "(UN x:C. B(x)) Int A  = (UN x:C. B(x) Int A)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   172
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   173
qed "UN_Int_distrib";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   174
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   175
val UN_simps = [UN_cons, Un_UN_distrib, UN_Un_distrib2, UN_Un_distrib, UN_Int_distrib, UN_0];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   176
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   177
(** Needed in State theory for the current definition of variables 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   178
where they are indexed by lists  **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   179
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   180
Goal "i:list(nat) ==> i:univ(0)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   181
by (dres_inst_tac [("B", "0")] list_into_univ 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   182
by (blast_tac (claset() addIs [nat_into_univ]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   183
by (assume_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   184
qed "list_nat_into_univ";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   185
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   186
(** To be moved to Update theory **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   187
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   188
Goalw [update_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   189
  "[| f:Pi(A,B); x:A; y:B(x) |] ==> f(x:=y):Pi(A, B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   190
by (asm_simp_tac (simpset() addsimps [domain_of_fun, cons_absorb, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   191
                                      apply_funtype, lam_type]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   192
qed "update_type2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   193
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   194
(** Simplication rules for Collect; To be moved elsewhere **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   195
Goal "{x:A. P(x)} Int A = {x:A. P(x)}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   196
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   197
qed "Collect_Int2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   198
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   199
Goal "A Int {x:A. P(x)} = {x:A. P(x)}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   200
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   201
qed "Collect_Int3";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   202
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   203
AddIffs [Collect_Int2, Collect_Int3];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   204
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   205
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   206
Goal "{x:A. P(x) | Q(x)} = Collect(A, P) Un Collect(A, Q)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   207
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   208
qed "Collect_disj_eq";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   209
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   210
Goal "{x:A. P(x) & Q(x)} = Collect(A, P) Int Collect(A, Q)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   211
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   212
qed "Collect_conj_eq";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   213
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   214
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   215