src/ZF/subset.ML
author paulson
Thu, 13 Jul 2000 13:04:48 +0200
changeset 9304 342cbcb9c0d8
parent 9211 6236c5285bd8
child 9907 473a6604da94
permissions -rw-r--r--
added an important default rule
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     1
(*  Title:      ZF/subset
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Derived rules involving subsets
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
Union and Intersection as lattice operations
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
(*** cons ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    12
Goal "[| a:C; B<=C |] ==> cons(a,B) <= C";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    13
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    14
qed "cons_subsetI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    16
Goal "B <= cons(a,B)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    17
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    18
qed "subset_consI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    20
Goal "cons(a,B)<=C <-> a:C & B<=C";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    21
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    22
qed "cons_subset_iff";
9304
342cbcb9c0d8 added an important default rule
paulson
parents: 9211
diff changeset
    23
AddIffs [cons_subset_iff];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
(*A safe special case of subset elimination, adding no new variables 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
  [| cons(a,B) <= C; [| a : C; B <= C |] ==> R |] ==> R *)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
    27
bind_thm ("cons_subsetE", cons_subset_iff RS iffD1 RS conjE);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    29
Goal "A<=0 <-> A=0";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    30
by (Blast_tac 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    31
qed "subset_empty_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    33
Goal "C<=cons(a,B) <-> C<=B | (a:C & C-{a} <= B)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    34
by (Blast_tac 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    35
qed "subset_cons_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
(*** succ ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    39
Goal "i <= succ(i)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    40
by (Blast_tac 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    41
qed "subset_succI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
(*But if j is an ordinal or is transitive, then i:j implies i<=j! 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    44
  See ordinal/Ord_succ_subsetI*)
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    45
Goalw [succ_def] "[| i:j;  i<=j |] ==> succ(i)<=j";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    46
by (Blast_tac 1) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    47
qed "succ_subsetI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    49
val major::prems= Goalw [succ_def]  
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
    "[| succ(i) <= j;  [| i:j;  i<=j |] ==> P \
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    51
\    |] ==> P";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    52
by (rtac (major RS cons_subsetE) 1);
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    53
by (REPEAT (ares_tac prems 1)) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9180
diff changeset
    54
qed "succ_subsetE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    55
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
(*** singletons ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    58
Goal "a:C ==> {a} <= C";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    59
by (Blast_tac 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    60
qed "singleton_subsetI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    62
Goal "{a} <= C  ==>  a:C";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    63
by (Blast_tac 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    64
qed "singleton_subsetD";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
(*** Big Union -- least upper bound of a set  ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    68
Goal "Union(A) <= C <-> (ALL x:A. x <= C)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    69
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    70
qed "Union_subset_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    72
Goal "B:A ==> B <= Union(A)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    73
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    74
qed "Union_upper";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    76
val [prem]= Goal
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    77
    "[| !!x. x:A ==> x<=C |] ==> Union(A) <= C";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    78
by (rtac (ballI RS (Union_subset_iff RS iffD2)) 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    79
by (etac prem 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    80
qed "Union_least";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
(*** Union of a family of sets  ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    83
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 4091
diff changeset
    84
Goal "A <= (UN i:I. B(i)) <-> A = (UN i:I. A Int B(i))";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    85
by (blast_tac (claset() addSEs [equalityE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 689
diff changeset
    86
qed "subset_UN_iff_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    88
Goal "(UN x:A. B(x)) <= C <-> (ALL x:A. B(x) <= C)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    89
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    90
qed "UN_subset_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    92
Goal "x:A ==> B(x) <= (UN x:A. B(x))";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    93
by (etac (RepFunI RS Union_upper) 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    94
qed "UN_upper";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    96
val [prem]= Goal
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    97
    "[| !!x. x:A ==> B(x)<=C |] ==> (UN x:A. B(x)) <= C";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    98
by (rtac (ballI RS (UN_subset_iff RS iffD2)) 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
    99
by (etac prem 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   100
qed "UN_least";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
(*** Big Intersection -- greatest lower bound of a nonempty set ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   105
Goal "a: A  ==>  C <= Inter(A) <-> (ALL x:A. C <= x)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   106
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   107
qed "Inter_subset_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   108
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   109
Goal "B:A ==> Inter(A) <= B";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   110
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   111
qed "Inter_lower";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   112
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   113
val [prem1,prem2]= Goal
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   114
    "[| a:A;  !!x. x:A ==> C<=x |] ==> C <= Inter(A)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   115
by (rtac ([prem1, ballI] MRS (Inter_subset_iff RS iffD2)) 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   116
by (etac prem2 1) ;
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   117
qed "Inter_greatest";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   118
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
(*** Intersection of a family of sets  ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   120
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   121
Goal "x:A ==> (INT x:A. B(x)) <= B(x)";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   122
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   123
qed "INT_lower";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   125
val [nonempty,prem] = Goal
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   126
    "[| a:A;  !!x. x:A ==> C<=B(x) |] ==> C <= (INT x:A. B(x))";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   127
by (rtac (nonempty RS RepFunI RS Inter_greatest) 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   128
by (REPEAT (eresolve_tac [RepFunE, prem, ssubst] 1));
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   129
qed "INT_greatest";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   131
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
(*** Finite Union -- the least upper bound of 2 sets ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   133
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   134
Goal "A Un B <= C <-> A <= C & B <= C";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   135
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   136
qed "Un_subset_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   137
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   138
Goal "A <= A Un B";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   139
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   140
qed "Un_upper1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   141
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   142
Goal "B <= A Un B";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   143
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   144
qed "Un_upper2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   145
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   146
Goal "[| A<=C;  B<=C |] ==> A Un B <= C";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   147
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   148
qed "Un_least";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   149
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   150
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   151
(*** Finite Intersection -- the greatest lower bound of 2 sets *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   152
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   153
Goal "C <= A Int B <-> C <= A & C <= B";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   154
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   155
qed "Int_subset_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   157
Goal "A Int B <= A";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   158
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   159
qed "Int_lower1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   160
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   161
Goal "A Int B <= B";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   162
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   163
qed "Int_lower2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   165
Goal "[| C<=A;  C<=B |] ==> C <= A Int B";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   166
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   167
qed "Int_greatest";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   168
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   169
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   170
(*** Set difference *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   171
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   172
Goal "A-B <= A";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   173
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   174
qed "Diff_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   176
Goal "[| C<=A;  C Int B = 0 |] ==> C <= A-B";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   177
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   178
qed "Diff_contains";
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 4091
diff changeset
   179
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 4091
diff changeset
   180
Goal "B <= A - cons(c,C)  <->  B<=A-C & c ~: B";
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 4091
diff changeset
   181
by (Blast_tac 1);
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 4091
diff changeset
   182
qed "subset_Diff_cons_iff";
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 4091
diff changeset
   183
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   184
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   185
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   186
(** Collect **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   187
9180
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   188
Goal "Collect(A,P) <= A";
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   189
by (Blast_tac 1);
3bda56c0d70d tidying and unbatchifying
paulson
parents: 5325
diff changeset
   190
qed "Collect_subset";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   191
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   192
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   193
(** RepFun **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   194
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 4091
diff changeset
   195
val prems = Goal "[| !!x. x:A ==> f(x): B |] ==> {f(x). x:A} <= B";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   196
by (blast_tac (claset() addIs prems) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 689
diff changeset
   197
qed "RepFun_subset";
689
bf95dada47ac ZF/Fixedpt/subset_cs: moved to ZF/subset.ML
lcp
parents: 0
diff changeset
   198
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   199
val subset_SIs =
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   200
    [subset_refl, cons_subsetI, subset_consI, 
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   201
     Union_least, UN_least, Un_least, 
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   202
     Inter_greatest, Int_greatest, RepFun_subset,
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   203
     Un_upper1, Un_upper2, Int_lower1, Int_lower2];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   204
689
bf95dada47ac ZF/Fixedpt/subset_cs: moved to ZF/subset.ML
lcp
parents: 0
diff changeset
   205
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   206
(*A claset for subset reasoning*)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   207
val subset_cs = claset() 
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   208
    delrules [subsetI, subsetCE]
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   209
    addSIs subset_SIs
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   210
    addIs  [Union_upper, Inter_lower]
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   211
    addSEs [cons_subsetE];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1745
diff changeset
   212