src/HOL/Set.ML
author paulson
Fri, 05 Nov 1999 12:47:29 +0100
changeset 8001 14c8843cd35b
parent 7969 7a20317850ab
child 8005 b64d86018785
permissions -rw-r--r--
new psubset lemma
clasohm@1465
     1
(*  Title:      HOL/set
clasohm@923
     2
    ID:         $Id$
clasohm@1465
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
clasohm@923
     4
    Copyright   1991  University of Cambridge
clasohm@923
     5
paulson@1985
     6
Set theory for higher-order logic.  A set is simply a predicate.
clasohm@923
     7
*)
clasohm@923
     8
nipkow@1548
     9
section "Relating predicates and sets";
nipkow@1548
    10
paulson@3469
    11
Addsimps [Collect_mem_eq];
paulson@3469
    12
AddIffs  [mem_Collect_eq];
paulson@2499
    13
paulson@5143
    14
Goal "P(a) ==> a : {x. P(x)}";
paulson@2499
    15
by (Asm_simp_tac 1);
clasohm@923
    16
qed "CollectI";
clasohm@923
    17
paulson@5316
    18
Goal "a : {x. P(x)} ==> P(a)";
paulson@2499
    19
by (Asm_full_simp_tac 1);
clasohm@923
    20
qed "CollectD";
clasohm@923
    21
wenzelm@7658
    22
bind_thm ("CollectE", make_elim CollectD);
wenzelm@7658
    23
paulson@5316
    24
val [prem] = Goal "[| !!x. (x:A) = (x:B) |] ==> A = B";
clasohm@923
    25
by (rtac (prem RS ext RS arg_cong RS box_equals) 1);
clasohm@923
    26
by (rtac Collect_mem_eq 1);
clasohm@923
    27
by (rtac Collect_mem_eq 1);
clasohm@923
    28
qed "set_ext";
clasohm@923
    29
paulson@5316
    30
val [prem] = Goal "[| !!x. P(x)=Q(x) |] ==> {x. P(x)} = {x. Q(x)}";
clasohm@923
    31
by (rtac (prem RS ext RS arg_cong) 1);
clasohm@923
    32
qed "Collect_cong";
clasohm@923
    33
clasohm@923
    34
val CollectE = make_elim CollectD;
clasohm@923
    35
paulson@2499
    36
AddSIs [CollectI];
paulson@2499
    37
AddSEs [CollectE];
paulson@2499
    38
paulson@2499
    39
nipkow@1548
    40
section "Bounded quantifiers";
clasohm@923
    41
paulson@5316
    42
val prems = Goalw [Ball_def]
clasohm@923
    43
    "[| !!x. x:A ==> P(x) |] ==> ! x:A. P(x)";
clasohm@923
    44
by (REPEAT (ares_tac (prems @ [allI,impI]) 1));
clasohm@923
    45
qed "ballI";
clasohm@923
    46
paulson@5316
    47
Goalw [Ball_def] "[| ! x:A. P(x);  x:A |] ==> P(x)";
paulson@5316
    48
by (Blast_tac 1);
clasohm@923
    49
qed "bspec";
clasohm@923
    50
paulson@5316
    51
val major::prems = Goalw [Ball_def]
clasohm@923
    52
    "[| ! x:A. P(x);  P(x) ==> Q;  x~:A ==> Q |] ==> Q";
clasohm@923
    53
by (rtac (major RS spec RS impCE) 1);
clasohm@923
    54
by (REPEAT (eresolve_tac prems 1));
clasohm@923
    55
qed "ballE";
clasohm@923
    56
clasohm@923
    57
(*Takes assumptions ! x:A.P(x) and a:A; creates assumption P(a)*)
clasohm@923
    58
fun ball_tac i = etac ballE i THEN contr_tac (i+1);
clasohm@923
    59
paulson@2499
    60
AddSIs [ballI];
paulson@2499
    61
AddEs  [ballE];
wenzelm@7441
    62
AddXDs [bspec];
oheimb@5521
    63
(* gives better instantiation for bound: *)
oheimb@5521
    64
claset_ref() := claset() addWrapper ("bspec", fn tac2 =>
oheimb@5521
    65
			 (dtac bspec THEN' atac) APPEND' tac2);
paulson@2499
    66
paulson@6006
    67
(*Normally the best argument order: P(x) constrains the choice of x:A*)
paulson@5316
    68
Goalw [Bex_def] "[| P(x);  x:A |] ==> ? x:A. P(x)";
paulson@5316
    69
by (Blast_tac 1);
clasohm@923
    70
qed "bexI";
clasohm@923
    71
paulson@6006
    72
(*The best argument order when there is only one x:A*)
paulson@6006
    73
Goalw [Bex_def] "[| x:A;  P(x) |] ==> ? x:A. P(x)";
paulson@6006
    74
by (Blast_tac 1);
paulson@6006
    75
qed "rev_bexI";
paulson@6006
    76
paulson@7031
    77
val prems = Goal 
paulson@7007
    78
   "[| ! x:A. ~P(x) ==> P(a);  a:A |] ==> ? x:A. P(x)";
paulson@7007
    79
by (rtac classical 1);
paulson@7007
    80
by (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1))  ;
paulson@7007
    81
qed "bexCI";
clasohm@923
    82
paulson@5316
    83
val major::prems = Goalw [Bex_def]
clasohm@923
    84
    "[| ? x:A. P(x);  !!x. [| x:A; P(x) |] ==> Q  |] ==> Q";
clasohm@923
    85
by (rtac (major RS exE) 1);
clasohm@923
    86
by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1));
clasohm@923
    87
qed "bexE";
clasohm@923
    88
paulson@2499
    89
AddIs  [bexI];
paulson@2499
    90
AddSEs [bexE];
paulson@2499
    91
paulson@3420
    92
(*Trival rewrite rule*)
wenzelm@5069
    93
Goal "(! x:A. P) = ((? x. x:A) --> P)";
wenzelm@4089
    94
by (simp_tac (simpset() addsimps [Ball_def]) 1);
paulson@3420
    95
qed "ball_triv";
paulson@1816
    96
paulson@1882
    97
(*Dual form for existentials*)
wenzelm@5069
    98
Goal "(? x:A. P) = ((? x. x:A) & P)";
wenzelm@4089
    99
by (simp_tac (simpset() addsimps [Bex_def]) 1);
paulson@3420
   100
qed "bex_triv";
paulson@1882
   101
paulson@3420
   102
Addsimps [ball_triv, bex_triv];
clasohm@923
   103
clasohm@923
   104
(** Congruence rules **)
clasohm@923
   105
paulson@6291
   106
val prems = Goalw [Ball_def]
clasohm@923
   107
    "[| A=B;  !!x. x:B ==> P(x) = Q(x) |] ==> \
clasohm@923
   108
\    (! x:A. P(x)) = (! x:B. Q(x))";
paulson@6291
   109
by (asm_simp_tac (simpset() addsimps prems) 1);
clasohm@923
   110
qed "ball_cong";
clasohm@923
   111
paulson@6291
   112
val prems = Goalw [Bex_def]
clasohm@923
   113
    "[| A=B;  !!x. x:B ==> P(x) = Q(x) |] ==> \
clasohm@923
   114
\    (? x:A. P(x)) = (? x:B. Q(x))";
paulson@6291
   115
by (asm_simp_tac (simpset() addcongs [conj_cong] addsimps prems) 1);
clasohm@923
   116
qed "bex_cong";
clasohm@923
   117
paulson@6291
   118
Addcongs [ball_cong,bex_cong];
paulson@6291
   119
nipkow@1548
   120
section "Subsets";
clasohm@923
   121
paulson@5316
   122
val prems = Goalw [subset_def] "(!!x. x:A ==> x:B) ==> A <= B";
clasohm@923
   123
by (REPEAT (ares_tac (prems @ [ballI]) 1));
clasohm@923
   124
qed "subsetI";
clasohm@923
   125
paulson@5649
   126
(*Map the type ('a set => anything) to just 'a.
paulson@5649
   127
  For overloading constants whose first argument has type "'a set" *)
paulson@5649
   128
fun overload_1st_set s = Blast.overloaded (s, HOLogic.dest_setT o domain_type);
paulson@5649
   129
paulson@4059
   130
(*While (:) is not, its type must be kept
paulson@4059
   131
  for overloading of = to work.*)
paulson@4240
   132
Blast.overloaded ("op :", domain_type);
paulson@5649
   133
paulson@5649
   134
overload_1st_set "Ball";		(*need UNION, INTER also?*)
paulson@5649
   135
overload_1st_set "Bex";
paulson@4059
   136
paulson@4469
   137
(*Image: retain the type of the set being expressed*)
paulson@5336
   138
Blast.overloaded ("op ``", domain_type);
paulson@2881
   139
clasohm@923
   140
(*Rule in Modus Ponens style*)
paulson@5316
   141
Goalw [subset_def] "[| A <= B;  c:A |] ==> c:B";
paulson@5316
   142
by (Blast_tac 1);
clasohm@923
   143
qed "subsetD";
wenzelm@7658
   144
AddXIs [subsetD];
clasohm@923
   145
clasohm@923
   146
(*The same, with reversed premises for use with etac -- cf rev_mp*)
paulson@7007
   147
Goal "[| c:A;  A <= B |] ==> c:B";
paulson@7007
   148
by (REPEAT (ares_tac [subsetD] 1)) ;
paulson@7007
   149
qed "rev_subsetD";
wenzelm@7658
   150
AddXIs [rev_subsetD];
clasohm@923
   151
paulson@1920
   152
(*Converts A<=B to x:A ==> x:B*)
paulson@1920
   153
fun impOfSubs th = th RSN (2, rev_subsetD);
paulson@1920
   154
paulson@7007
   155
Goal "[| A <= B; c ~: B |] ==> c ~: A";
paulson@7007
   156
by (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ;
paulson@7007
   157
qed "contra_subsetD";
paulson@1841
   158
paulson@7007
   159
Goal "[| c ~: B;  A <= B |] ==> c ~: A";
paulson@7007
   160
by (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ;
paulson@7007
   161
qed "rev_contra_subsetD";
paulson@1841
   162
clasohm@923
   163
(*Classical elimination rule*)
paulson@5316
   164
val major::prems = Goalw [subset_def] 
clasohm@923
   165
    "[| A <= B;  c~:A ==> P;  c:B ==> P |] ==> P";
clasohm@923
   166
by (rtac (major RS ballE) 1);
clasohm@923
   167
by (REPEAT (eresolve_tac prems 1));
clasohm@923
   168
qed "subsetCE";
clasohm@923
   169
clasohm@923
   170
(*Takes assumptions A<=B; c:A and creates the assumption c:B *)
clasohm@923
   171
fun set_mp_tac i = etac subsetCE i  THEN  mp_tac i;
clasohm@923
   172
paulson@2499
   173
AddSIs [subsetI];
paulson@2499
   174
AddEs  [subsetD, subsetCE];
paulson@2499
   175
paulson@7007
   176
Goal "A <= (A::'a set)";
paulson@7007
   177
by (Fast_tac 1);
paulson@7007
   178
qed "subset_refl";		(*Blast_tac would try order_refl and fail*)
clasohm@923
   179
paulson@5316
   180
Goal "[| A<=B;  B<=C |] ==> A<=(C::'a set)";
paulson@2891
   181
by (Blast_tac 1);
clasohm@923
   182
qed "subset_trans";
clasohm@923
   183
clasohm@923
   184
nipkow@1548
   185
section "Equality";
clasohm@923
   186
clasohm@923
   187
(*Anti-symmetry of the subset relation*)
paulson@5316
   188
Goal "[| A <= B;  B <= A |] ==> A = (B::'a set)";
paulson@5318
   189
by (rtac set_ext 1);
paulson@5316
   190
by (blast_tac (claset() addIs [subsetD]) 1);
clasohm@923
   191
qed "subset_antisym";
clasohm@923
   192
val equalityI = subset_antisym;
clasohm@923
   193
berghofe@1762
   194
AddSIs [equalityI];
berghofe@1762
   195
clasohm@923
   196
(* Equality rules from ZF set theory -- are they appropriate here? *)
paulson@5316
   197
Goal "A = B ==> A<=(B::'a set)";
paulson@5316
   198
by (etac ssubst 1);
clasohm@923
   199
by (rtac subset_refl 1);
clasohm@923
   200
qed "equalityD1";
clasohm@923
   201
paulson@5316
   202
Goal "A = B ==> B<=(A::'a set)";
paulson@5316
   203
by (etac ssubst 1);
clasohm@923
   204
by (rtac subset_refl 1);
clasohm@923
   205
qed "equalityD2";
clasohm@923
   206
paulson@5316
   207
val prems = Goal
clasohm@923
   208
    "[| A = B;  [| A<=B; B<=(A::'a set) |] ==> P |]  ==>  P";
clasohm@923
   209
by (resolve_tac prems 1);
clasohm@923
   210
by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1));
clasohm@923
   211
qed "equalityE";
clasohm@923
   212
paulson@5316
   213
val major::prems = Goal
clasohm@923
   214
    "[| A = B;  [| c:A; c:B |] ==> P;  [| c~:A; c~:B |] ==> P |]  ==>  P";
clasohm@923
   215
by (rtac (major RS equalityE) 1);
clasohm@923
   216
by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1));
clasohm@923
   217
qed "equalityCE";
clasohm@923
   218
clasohm@923
   219
(*Lemma for creating induction formulae -- for "pattern matching" on p
clasohm@923
   220
  To make the induction hypotheses usable, apply "spec" or "bspec" to
clasohm@923
   221
  put universal quantifiers over the free variables in p. *)
paulson@5316
   222
val prems = Goal 
clasohm@923
   223
    "[| p:A;  !!z. z:A ==> p=z --> R |] ==> R";
clasohm@923
   224
by (rtac mp 1);
clasohm@923
   225
by (REPEAT (resolve_tac (refl::prems) 1));
clasohm@923
   226
qed "setup_induction";
clasohm@923
   227
clasohm@923
   228
paulson@4159
   229
section "The universal set -- UNIV";
paulson@4159
   230
paulson@7031
   231
Goalw [UNIV_def] "x : UNIV";
paulson@7031
   232
by (rtac CollectI 1);
paulson@7031
   233
by (rtac TrueI 1);
paulson@7031
   234
qed "UNIV_I";
paulson@4159
   235
paulson@4434
   236
Addsimps [UNIV_I];
paulson@4434
   237
AddIs    [UNIV_I];  (*unsafe makes it less likely to cause problems*)
paulson@4159
   238
paulson@7031
   239
Goal "A <= UNIV";
paulson@7031
   240
by (rtac subsetI 1);
paulson@7031
   241
by (rtac UNIV_I 1);
paulson@7031
   242
qed "subset_UNIV";
paulson@4159
   243
paulson@4159
   244
(** Eta-contracting these two rules (to remove P) causes them to be ignored
paulson@4159
   245
    because of their interaction with congruence rules. **)
paulson@4159
   246
wenzelm@5069
   247
Goalw [Ball_def] "Ball UNIV P = All P";
paulson@4159
   248
by (Simp_tac 1);
paulson@4159
   249
qed "ball_UNIV";
paulson@4159
   250
wenzelm@5069
   251
Goalw [Bex_def] "Bex UNIV P = Ex P";
paulson@4159
   252
by (Simp_tac 1);
paulson@4159
   253
qed "bex_UNIV";
paulson@4159
   254
Addsimps [ball_UNIV, bex_UNIV];
paulson@4159
   255
paulson@4159
   256
paulson@2858
   257
section "The empty set -- {}";
paulson@2858
   258
paulson@7007
   259
Goalw [empty_def] "(c : {}) = False";
paulson@7007
   260
by (Blast_tac 1) ;
paulson@7007
   261
qed "empty_iff";
paulson@2858
   262
paulson@2858
   263
Addsimps [empty_iff];
paulson@2858
   264
paulson@7007
   265
Goal "a:{} ==> P";
paulson@7007
   266
by (Full_simp_tac 1);
paulson@7007
   267
qed "emptyE";
paulson@2858
   268
paulson@2858
   269
AddSEs [emptyE];
paulson@2858
   270
paulson@7007
   271
Goal "{} <= A";
paulson@7007
   272
by (Blast_tac 1) ;
paulson@7007
   273
qed "empty_subsetI";
paulson@2858
   274
paulson@5256
   275
(*One effect is to delete the ASSUMPTION {} <= A*)
paulson@5256
   276
AddIffs [empty_subsetI];
paulson@5256
   277
paulson@7031
   278
val [prem]= Goal "[| !!y. y:A ==> False |] ==> A={}";
paulson@7007
   279
by (blast_tac (claset() addIs [prem RS FalseE]) 1) ;
paulson@7007
   280
qed "equals0I";
paulson@2858
   281
paulson@5256
   282
(*Use for reasoning about disjointness: A Int B = {} *)
paulson@7007
   283
Goal "A={} ==> a ~: A";
paulson@7007
   284
by (Blast_tac 1) ;
paulson@7007
   285
qed "equals0D";
paulson@2858
   286
paulson@5450
   287
AddDs [equals0D, sym RS equals0D];
paulson@5256
   288
wenzelm@5069
   289
Goalw [Ball_def] "Ball {} P = True";
paulson@4159
   290
by (Simp_tac 1);
paulson@4159
   291
qed "ball_empty";
paulson@4159
   292
wenzelm@5069
   293
Goalw [Bex_def] "Bex {} P = False";
paulson@4159
   294
by (Simp_tac 1);
paulson@4159
   295
qed "bex_empty";
paulson@4159
   296
Addsimps [ball_empty, bex_empty];
paulson@4159
   297
wenzelm@5069
   298
Goal "UNIV ~= {}";
paulson@4159
   299
by (blast_tac (claset() addEs [equalityE]) 1);
paulson@4159
   300
qed "UNIV_not_empty";
paulson@4159
   301
AddIffs [UNIV_not_empty];
paulson@4159
   302
paulson@4159
   303
paulson@2858
   304
paulson@2858
   305
section "The Powerset operator -- Pow";
paulson@2858
   306
paulson@7007
   307
Goalw [Pow_def] "(A : Pow(B)) = (A <= B)";
paulson@7007
   308
by (Asm_simp_tac 1);
paulson@7007
   309
qed "Pow_iff";
paulson@2858
   310
paulson@2858
   311
AddIffs [Pow_iff]; 
paulson@2858
   312
paulson@7031
   313
Goalw [Pow_def] "A <= B ==> A : Pow(B)";
paulson@7007
   314
by (etac CollectI 1);
paulson@7007
   315
qed "PowI";
paulson@2858
   316
paulson@7031
   317
Goalw [Pow_def] "A : Pow(B)  ==>  A<=B";
paulson@7007
   318
by (etac CollectD 1);
paulson@7007
   319
qed "PowD";
paulson@7007
   320
paulson@2858
   321
paulson@2858
   322
val Pow_bottom = empty_subsetI RS PowI;        (* {}: Pow(B) *)
paulson@2858
   323
val Pow_top = subset_refl RS PowI;             (* A : Pow(A) *)
paulson@2858
   324
paulson@2858
   325
paulson@5931
   326
section "Set complement";
clasohm@923
   327
paulson@7031
   328
Goalw [Compl_def] "(c : -A) = (c~:A)";
paulson@7031
   329
by (Blast_tac 1);
paulson@7031
   330
qed "Compl_iff";
paulson@2499
   331
paulson@2499
   332
Addsimps [Compl_iff];
paulson@2499
   333
paulson@5490
   334
val prems = Goalw [Compl_def] "[| c:A ==> False |] ==> c : -A";
clasohm@923
   335
by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1));
clasohm@923
   336
qed "ComplI";
clasohm@923
   337
clasohm@923
   338
(*This form, with negated conclusion, works well with the Classical prover.
clasohm@923
   339
  Negated assumptions behave like formulae on the right side of the notional
clasohm@923
   340
  turnstile...*)
paulson@5490
   341
Goalw [Compl_def] "c : -A ==> c~:A";
paulson@5316
   342
by (etac CollectD 1);
clasohm@923
   343
qed "ComplD";
clasohm@923
   344
clasohm@923
   345
val ComplE = make_elim ComplD;
clasohm@923
   346
paulson@2499
   347
AddSIs [ComplI];
paulson@2499
   348
AddSEs [ComplE];
paulson@1640
   349
clasohm@923
   350
nipkow@1548
   351
section "Binary union -- Un";
clasohm@923
   352
paulson@7031
   353
Goalw [Un_def] "(c : A Un B) = (c:A | c:B)";
paulson@7031
   354
by (Blast_tac 1);
paulson@7031
   355
qed "Un_iff";
paulson@2499
   356
Addsimps [Un_iff];
paulson@2499
   357
paulson@5143
   358
Goal "c:A ==> c : A Un B";
paulson@2499
   359
by (Asm_simp_tac 1);
clasohm@923
   360
qed "UnI1";
clasohm@923
   361
paulson@5143
   362
Goal "c:B ==> c : A Un B";
paulson@2499
   363
by (Asm_simp_tac 1);
clasohm@923
   364
qed "UnI2";
clasohm@923
   365
clasohm@923
   366
(*Classical introduction rule: no commitment to A vs B*)
paulson@7007
   367
paulson@7031
   368
val prems = Goal "(c~:B ==> c:A) ==> c : A Un B";
paulson@7007
   369
by (Simp_tac 1);
paulson@7007
   370
by (REPEAT (ares_tac (prems@[disjCI]) 1)) ;
paulson@7007
   371
qed "UnCI";
clasohm@923
   372
paulson@5316
   373
val major::prems = Goalw [Un_def]
clasohm@923
   374
    "[| c : A Un B;  c:A ==> P;  c:B ==> P |] ==> P";
clasohm@923
   375
by (rtac (major RS CollectD RS disjE) 1);
clasohm@923
   376
by (REPEAT (eresolve_tac prems 1));
clasohm@923
   377
qed "UnE";
clasohm@923
   378
paulson@2499
   379
AddSIs [UnCI];
paulson@2499
   380
AddSEs [UnE];
paulson@1640
   381
clasohm@923
   382
nipkow@1548
   383
section "Binary intersection -- Int";
clasohm@923
   384
paulson@7031
   385
Goalw [Int_def] "(c : A Int B) = (c:A & c:B)";
paulson@7031
   386
by (Blast_tac 1);
paulson@7031
   387
qed "Int_iff";
paulson@2499
   388
Addsimps [Int_iff];
paulson@2499
   389
paulson@5143
   390
Goal "[| c:A;  c:B |] ==> c : A Int B";
paulson@2499
   391
by (Asm_simp_tac 1);
clasohm@923
   392
qed "IntI";
clasohm@923
   393
paulson@5143
   394
Goal "c : A Int B ==> c:A";
paulson@2499
   395
by (Asm_full_simp_tac 1);
clasohm@923
   396
qed "IntD1";
clasohm@923
   397
paulson@5143
   398
Goal "c : A Int B ==> c:B";
paulson@2499
   399
by (Asm_full_simp_tac 1);
clasohm@923
   400
qed "IntD2";
clasohm@923
   401
paulson@5316
   402
val [major,minor] = Goal
clasohm@923
   403
    "[| c : A Int B;  [| c:A; c:B |] ==> P |] ==> P";
clasohm@923
   404
by (rtac minor 1);
clasohm@923
   405
by (rtac (major RS IntD1) 1);
clasohm@923
   406
by (rtac (major RS IntD2) 1);
clasohm@923
   407
qed "IntE";
clasohm@923
   408
paulson@2499
   409
AddSIs [IntI];
paulson@2499
   410
AddSEs [IntE];
clasohm@923
   411
nipkow@1548
   412
section "Set difference";
clasohm@923
   413
paulson@7031
   414
Goalw [set_diff_def] "(c : A-B) = (c:A & c~:B)";
paulson@7031
   415
by (Blast_tac 1);
paulson@7031
   416
qed "Diff_iff";
paulson@2499
   417
Addsimps [Diff_iff];
clasohm@923
   418
paulson@7007
   419
Goal "[| c : A;  c ~: B |] ==> c : A - B";
paulson@7007
   420
by (Asm_simp_tac 1) ;
paulson@7007
   421
qed "DiffI";
clasohm@923
   422
paulson@7007
   423
Goal "c : A - B ==> c : A";
paulson@7007
   424
by (Asm_full_simp_tac 1) ;
paulson@7007
   425
qed "DiffD1";
paulson@2499
   426
paulson@7007
   427
Goal "[| c : A - B;  c : B |] ==> P";
paulson@7007
   428
by (Asm_full_simp_tac 1) ;
paulson@7007
   429
qed "DiffD2";
paulson@2499
   430
paulson@7031
   431
val prems = Goal "[| c : A - B;  [| c:A; c~:B |] ==> P |] ==> P";
paulson@7007
   432
by (resolve_tac prems 1);
paulson@7007
   433
by (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ;
paulson@7007
   434
qed "DiffE";
clasohm@923
   435
paulson@2499
   436
AddSIs [DiffI];
paulson@2499
   437
AddSEs [DiffE];
clasohm@923
   438
clasohm@923
   439
nipkow@1548
   440
section "Augmenting a set -- insert";
clasohm@923
   441
paulson@7031
   442
Goalw [insert_def] "a : insert b A = (a=b | a:A)";
paulson@7031
   443
by (Blast_tac 1);
paulson@7031
   444
qed "insert_iff";
paulson@2499
   445
Addsimps [insert_iff];
paulson@2499
   446
paulson@7031
   447
Goal "a : insert a B";
paulson@7007
   448
by (Simp_tac 1);
paulson@7007
   449
qed "insertI1";
paulson@2499
   450
paulson@7007
   451
Goal "!!a. a : B ==> a : insert b B";
paulson@7007
   452
by (Asm_simp_tac 1);
paulson@7007
   453
qed "insertI2";
clasohm@923
   454
paulson@7007
   455
val major::prems = Goalw [insert_def]
paulson@7007
   456
    "[| a : insert b A;  a=b ==> P;  a:A ==> P |] ==> P";
paulson@7007
   457
by (rtac (major RS UnE) 1);
paulson@7007
   458
by (REPEAT (eresolve_tac (prems @ [CollectE]) 1));
paulson@7007
   459
qed "insertE";
clasohm@923
   460
clasohm@923
   461
(*Classical introduction rule*)
paulson@7031
   462
val prems = Goal "(a~:B ==> a=b) ==> a: insert b B";
paulson@7007
   463
by (Simp_tac 1);
paulson@7007
   464
by (REPEAT (ares_tac (prems@[disjCI]) 1)) ;
paulson@7007
   465
qed "insertCI";
paulson@2499
   466
paulson@2499
   467
AddSIs [insertCI]; 
paulson@2499
   468
AddSEs [insertE];
clasohm@923
   469
oheimb@7496
   470
Goal "A <= insert x B ==> A <= B & x ~: A | (? B'. A = insert x B' & B' <= B)";
oheimb@7496
   471
by (case_tac "x:A" 1);
oheimb@7496
   472
by  (Fast_tac 2);
wenzelm@7499
   473
by (rtac disjI2 1);
oheimb@7496
   474
by (res_inst_tac [("x","A-{x}")] exI 1);
oheimb@7496
   475
by (Fast_tac 1);
oheimb@7496
   476
qed "subset_insertD";
oheimb@7496
   477
nipkow@1548
   478
section "Singletons, using insert";
clasohm@923
   479
paulson@7007
   480
Goal "a : {a}";
paulson@7007
   481
by (rtac insertI1 1) ;
paulson@7007
   482
qed "singletonI";
clasohm@923
   483
paulson@5143
   484
Goal "b : {a} ==> b=a";
paulson@2891
   485
by (Blast_tac 1);
clasohm@923
   486
qed "singletonD";
clasohm@923
   487
oheimb@1776
   488
bind_thm ("singletonE", make_elim singletonD);
oheimb@1776
   489
paulson@7007
   490
Goal "(b : {a}) = (b=a)";
paulson@7007
   491
by (Blast_tac 1);
paulson@7007
   492
qed "singleton_iff";
clasohm@923
   493
paulson@5143
   494
Goal "{a}={b} ==> a=b";
wenzelm@4089
   495
by (blast_tac (claset() addEs [equalityE]) 1);
clasohm@923
   496
qed "singleton_inject";
clasohm@923
   497
paulson@2858
   498
(*Redundant? But unlike insertCI, it proves the subgoal immediately!*)
paulson@2858
   499
AddSIs [singletonI];   
paulson@2499
   500
AddSDs [singleton_inject];
paulson@3718
   501
AddSEs [singletonE];
paulson@2499
   502
oheimb@7969
   503
Goal "{b} = insert a A = (a = b & A <= {b})";
oheimb@7496
   504
by (safe_tac (claset() addSEs [equalityE]));
oheimb@7496
   505
by   (ALLGOALS Blast_tac);
oheimb@7496
   506
qed "singleton_insert_inj_eq";
oheimb@7496
   507
oheimb@7969
   508
Goal "(insert a A = {b} ) = (a = b & A <= {b})";
oheimb@7969
   509
by (rtac (singleton_insert_inj_eq RS (eq_sym_conv RS trans)) 1);
oheimb@7969
   510
qed "singleton_insert_inj_eq'";
oheimb@7969
   511
oheimb@7496
   512
Goal "A <= {x} ==> A={} | A = {x}";
oheimb@7496
   513
by (Fast_tac 1);
oheimb@7496
   514
qed "subset_singletonD";
oheimb@7496
   515
wenzelm@5069
   516
Goal "{x. x=a} = {a}";
wenzelm@4423
   517
by (Blast_tac 1);
nipkow@3582
   518
qed "singleton_conv";
nipkow@3582
   519
Addsimps [singleton_conv];
nipkow@1531
   520
nipkow@5600
   521
Goal "{x. a=x} = {a}";
paulson@6301
   522
by (Blast_tac 1);
nipkow@5600
   523
qed "singleton_conv2";
nipkow@5600
   524
Addsimps [singleton_conv2];
nipkow@5600
   525
nipkow@1531
   526
nipkow@1548
   527
section "Unions of families -- UNION x:A. B(x) is Union(B``A)";
clasohm@923
   528
wenzelm@5069
   529
Goalw [UNION_def] "(b: (UN x:A. B(x))) = (EX x:A. b: B(x))";
paulson@2891
   530
by (Blast_tac 1);
paulson@2499
   531
qed "UN_iff";
paulson@2499
   532
paulson@2499
   533
Addsimps [UN_iff];
paulson@2499
   534
clasohm@923
   535
(*The order of the premises presupposes that A is rigid; b may be flexible*)
paulson@5143
   536
Goal "[| a:A;  b: B(a) |] ==> b: (UN x:A. B(x))";
paulson@4477
   537
by Auto_tac;
clasohm@923
   538
qed "UN_I";
clasohm@923
   539
paulson@5316
   540
val major::prems = Goalw [UNION_def]
clasohm@923
   541
    "[| b : (UN x:A. B(x));  !!x.[| x:A;  b: B(x) |] ==> R |] ==> R";
clasohm@923
   542
by (rtac (major RS CollectD RS bexE) 1);
clasohm@923
   543
by (REPEAT (ares_tac prems 1));
clasohm@923
   544
qed "UN_E";
clasohm@923
   545
paulson@2499
   546
AddIs  [UN_I];
paulson@2499
   547
AddSEs [UN_E];
paulson@2499
   548
paulson@6291
   549
val prems = Goalw [UNION_def]
clasohm@923
   550
    "[| A=B;  !!x. x:B ==> C(x) = D(x) |] ==> \
clasohm@923
   551
\    (UN x:A. C(x)) = (UN x:B. D(x))";
paulson@6291
   552
by (asm_simp_tac (simpset() addsimps prems) 1);
clasohm@923
   553
qed "UN_cong";
clasohm@923
   554
clasohm@923
   555
nipkow@1548
   556
section "Intersections of families -- INTER x:A. B(x) is Inter(B``A)";
clasohm@923
   557
wenzelm@5069
   558
Goalw [INTER_def] "(b: (INT x:A. B(x))) = (ALL x:A. b: B(x))";
paulson@4477
   559
by Auto_tac;
paulson@2499
   560
qed "INT_iff";
paulson@2499
   561
paulson@2499
   562
Addsimps [INT_iff];
paulson@2499
   563
paulson@5316
   564
val prems = Goalw [INTER_def]
clasohm@923
   565
    "(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))";
clasohm@923
   566
by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1));
clasohm@923
   567
qed "INT_I";
clasohm@923
   568
paulson@5143
   569
Goal "[| b : (INT x:A. B(x));  a:A |] ==> b: B(a)";
paulson@4477
   570
by Auto_tac;
clasohm@923
   571
qed "INT_D";
clasohm@923
   572
clasohm@923
   573
(*"Classical" elimination -- by the Excluded Middle on a:A *)
paulson@5316
   574
val major::prems = Goalw [INTER_def]
clasohm@923
   575
    "[| b : (INT x:A. B(x));  b: B(a) ==> R;  a~:A ==> R |] ==> R";
clasohm@923
   576
by (rtac (major RS CollectD RS ballE) 1);
clasohm@923
   577
by (REPEAT (eresolve_tac prems 1));
clasohm@923
   578
qed "INT_E";
clasohm@923
   579
paulson@2499
   580
AddSIs [INT_I];
paulson@2499
   581
AddEs  [INT_D, INT_E];
paulson@2499
   582
paulson@6291
   583
val prems = Goalw [INTER_def]
clasohm@923
   584
    "[| A=B;  !!x. x:B ==> C(x) = D(x) |] ==> \
clasohm@923
   585
\    (INT x:A. C(x)) = (INT x:B. D(x))";
paulson@6291
   586
by (asm_simp_tac (simpset() addsimps prems) 1);
clasohm@923
   587
qed "INT_cong";
clasohm@923
   588
clasohm@923
   589
nipkow@1548
   590
section "Union";
clasohm@923
   591
wenzelm@5069
   592
Goalw [Union_def] "(A : Union(C)) = (EX X:C. A:X)";
paulson@2891
   593
by (Blast_tac 1);
paulson@2499
   594
qed "Union_iff";
paulson@2499
   595
paulson@2499
   596
Addsimps [Union_iff];
paulson@2499
   597
clasohm@923
   598
(*The order of the premises presupposes that C is rigid; A may be flexible*)
paulson@5143
   599
Goal "[| X:C;  A:X |] ==> A : Union(C)";
paulson@4477
   600
by Auto_tac;
clasohm@923
   601
qed "UnionI";
clasohm@923
   602
paulson@5316
   603
val major::prems = Goalw [Union_def]
clasohm@923
   604
    "[| A : Union(C);  !!X.[| A:X;  X:C |] ==> R |] ==> R";
clasohm@923
   605
by (rtac (major RS UN_E) 1);
clasohm@923
   606
by (REPEAT (ares_tac prems 1));
clasohm@923
   607
qed "UnionE";
clasohm@923
   608
paulson@2499
   609
AddIs  [UnionI];
paulson@2499
   610
AddSEs [UnionE];
paulson@2499
   611
paulson@2499
   612
nipkow@1548
   613
section "Inter";
clasohm@923
   614
wenzelm@5069
   615
Goalw [Inter_def] "(A : Inter(C)) = (ALL X:C. A:X)";
paulson@2891
   616
by (Blast_tac 1);
paulson@2499
   617
qed "Inter_iff";
paulson@2499
   618
paulson@2499
   619
Addsimps [Inter_iff];
paulson@2499
   620
paulson@5316
   621
val prems = Goalw [Inter_def]
clasohm@923
   622
    "[| !!X. X:C ==> A:X |] ==> A : Inter(C)";
clasohm@923
   623
by (REPEAT (ares_tac ([INT_I] @ prems) 1));
clasohm@923
   624
qed "InterI";
clasohm@923
   625
clasohm@923
   626
(*A "destruct" rule -- every X in C contains A as an element, but
clasohm@923
   627
  A:X can hold when X:C does not!  This rule is analogous to "spec". *)
paulson@5143
   628
Goal "[| A : Inter(C);  X:C |] ==> A:X";
paulson@4477
   629
by Auto_tac;
clasohm@923
   630
qed "InterD";
clasohm@923
   631
clasohm@923
   632
(*"Classical" elimination rule -- does not require proving X:C *)
paulson@5316
   633
val major::prems = Goalw [Inter_def]
paulson@2721
   634
    "[| A : Inter(C);  X~:C ==> R;  A:X ==> R |] ==> R";
clasohm@923
   635
by (rtac (major RS INT_E) 1);
clasohm@923
   636
by (REPEAT (eresolve_tac prems 1));
clasohm@923
   637
qed "InterE";
clasohm@923
   638
paulson@2499
   639
AddSIs [InterI];
paulson@2499
   640
AddEs  [InterD, InterE];
paulson@2499
   641
paulson@2499
   642
nipkow@2912
   643
(*** Image of a set under a function ***)
nipkow@2912
   644
nipkow@2912
   645
(*Frequently b does not have the syntactic form of f(x).*)
paulson@5316
   646
Goalw [image_def] "[| b=f(x);  x:A |] ==> b : f``A";
paulson@5316
   647
by (Blast_tac 1);
nipkow@2912
   648
qed "image_eqI";
nipkow@3909
   649
Addsimps [image_eqI];
nipkow@2912
   650
nipkow@2912
   651
bind_thm ("imageI", refl RS image_eqI);
nipkow@2912
   652
nipkow@2912
   653
(*The eta-expansion gives variable-name preservation.*)
paulson@5316
   654
val major::prems = Goalw [image_def]
wenzelm@3842
   655
    "[| b : (%x. f(x))``A;  !!x.[| b=f(x);  x:A |] ==> P |] ==> P"; 
nipkow@2912
   656
by (rtac (major RS CollectD RS bexE) 1);
nipkow@2912
   657
by (REPEAT (ares_tac prems 1));
nipkow@2912
   658
qed "imageE";
nipkow@2912
   659
nipkow@2912
   660
AddIs  [image_eqI];
nipkow@2912
   661
AddSEs [imageE]; 
nipkow@2912
   662
wenzelm@5069
   663
Goal "f``(A Un B) = f``A Un f``B";
paulson@2935
   664
by (Blast_tac 1);
nipkow@2912
   665
qed "image_Un";
nipkow@2912
   666
wenzelm@5069
   667
Goal "(z : f``A) = (EX x:A. z = f x)";
paulson@3960
   668
by (Blast_tac 1);
paulson@3960
   669
qed "image_iff";
paulson@3960
   670
paulson@4523
   671
(*This rewrite rule would confuse users if made default.*)
wenzelm@5069
   672
Goal "(f``A <= B) = (ALL x:A. f(x): B)";
paulson@4523
   673
by (Blast_tac 1);
paulson@4523
   674
qed "image_subset_iff";
paulson@4523
   675
paulson@4523
   676
(*Replaces the three steps subsetI, imageE, hyp_subst_tac, but breaks too
paulson@4523
   677
  many existing proofs.*)
paulson@5316
   678
val prems = Goal "(!!x. x:A ==> f(x) : B) ==> f``A <= B";
paulson@4510
   679
by (blast_tac (claset() addIs prems) 1);
paulson@4510
   680
qed "image_subsetI";
paulson@4510
   681
nipkow@2912
   682
nipkow@2912
   683
(*** Range of a function -- just a translation for image! ***)
nipkow@2912
   684
paulson@5143
   685
Goal "b=f(x) ==> b : range(f)";
nipkow@2912
   686
by (EVERY1 [etac image_eqI, rtac UNIV_I]);
nipkow@2912
   687
bind_thm ("range_eqI", UNIV_I RSN (2,image_eqI));
nipkow@2912
   688
nipkow@2912
   689
bind_thm ("rangeI", UNIV_I RS imageI);
nipkow@2912
   690
paulson@5316
   691
val [major,minor] = Goal 
wenzelm@3842
   692
    "[| b : range(%x. f(x));  !!x. b=f(x) ==> P |] ==> P"; 
nipkow@2912
   693
by (rtac (major RS imageE) 1);
nipkow@2912
   694
by (etac minor 1);
nipkow@2912
   695
qed "rangeE";
nipkow@2912
   696
oheimb@1776
   697
oheimb@1776
   698
(*** Set reasoning tools ***)
oheimb@1776
   699
oheimb@1776
   700
paulson@3912
   701
(** Rewrite rules for boolean case-splitting: faster than 
nipkow@4830
   702
	addsplits[split_if]
paulson@3912
   703
**)
paulson@3912
   704
nipkow@4830
   705
bind_thm ("split_if_eq1", read_instantiate [("P", "%x. x = ?b")] split_if);
nipkow@4830
   706
bind_thm ("split_if_eq2", read_instantiate [("P", "%x. ?a = x")] split_if);
paulson@3912
   707
paulson@5237
   708
(*Split ifs on either side of the membership relation.
paulson@5237
   709
	Not for Addsimps -- can cause goals to blow up!*)
nipkow@4830
   710
bind_thm ("split_if_mem1", 
wenzelm@6394
   711
    read_instantiate_sg (Theory.sign_of Set.thy) [("P", "%x. x : ?b")] split_if);
nipkow@4830
   712
bind_thm ("split_if_mem2", 
wenzelm@6394
   713
    read_instantiate_sg (Theory.sign_of Set.thy) [("P", "%x. ?a : x")] split_if);
paulson@3912
   714
nipkow@4830
   715
val split_ifs = [if_bool_eq_conj, split_if_eq1, split_if_eq2,
nipkow@4830
   716
		  split_if_mem1, split_if_mem2];
paulson@3912
   717
paulson@3912
   718
wenzelm@4089
   719
(*Each of these has ALREADY been added to simpset() above.*)
paulson@2024
   720
val mem_simps = [insert_iff, empty_iff, Un_iff, Int_iff, Compl_iff, Diff_iff, 
paulson@4159
   721
                 mem_Collect_eq, UN_iff, Union_iff, INT_iff, Inter_iff];
oheimb@1776
   722
oheimb@1776
   723
val mksimps_pairs = ("Ball",[bspec]) :: mksimps_pairs;
oheimb@1776
   724
paulson@6291
   725
simpset_ref() := simpset() setmksimps (mksimps mksimps_pairs);
nipkow@3222
   726
paulson@5256
   727
Addsimps[subset_UNIV, subset_refl];
nipkow@3222
   728
nipkow@3222
   729
paulson@8001
   730
(*** The 'proper subset' relation (<) ***)
nipkow@3222
   731
wenzelm@5069
   732
Goalw [psubset_def] "!!A::'a set. [| A <= B; A ~= B |] ==> A<B";
nipkow@3222
   733
by (Blast_tac 1);
nipkow@3222
   734
qed "psubsetI";
wenzelm@7658
   735
AddXIs [psubsetI];
nipkow@3222
   736
paulson@5148
   737
Goalw [psubset_def] "A < insert x B ==> (x ~: A) & A<=B | x:A & A-{x}<B";
paulson@4477
   738
by Auto_tac;
nipkow@3222
   739
qed "psubset_insertD";
paulson@4059
   740
paulson@4059
   741
bind_thm ("psubset_eq", psubset_def RS meta_eq_to_obj_eq);
wenzelm@6443
   742
wenzelm@6443
   743
bind_thm ("psubset_imp_subset", psubset_eq RS iffD1 RS conjunct1);
wenzelm@6443
   744
wenzelm@6443
   745
Goal"[| (A::'a set) < B; B <= C |] ==> A < C";
wenzelm@6443
   746
by (auto_tac (claset(), simpset() addsimps [psubset_eq]));
wenzelm@6443
   747
qed "psubset_subset_trans";
wenzelm@6443
   748
wenzelm@6443
   749
Goal"[| (A::'a set) <= B; B < C|] ==> A < C";
wenzelm@6443
   750
by (auto_tac (claset(), simpset() addsimps [psubset_eq]));
wenzelm@6443
   751
qed "subset_psubset_trans";
berghofe@7717
   752
paulson@8001
   753
Goalw [psubset_def] "A < B ==> EX b. b : (B - A)";
paulson@8001
   754
by (Blast_tac 1);
paulson@8001
   755
qed "psubset_imp_ex_mem";
paulson@8001
   756
berghofe@7717
   757
berghofe@7717
   758
(* attributes *)
berghofe@7717
   759
berghofe@7717
   760
local
berghofe@7717
   761
berghofe@7717
   762
fun gen_rulify_prems x =
berghofe@7717
   763
  Attrib.no_args (Drule.rule_attribute (fn _ => (standard o
berghofe@7717
   764
    rule_by_tactic (REPEAT (ALLGOALS (resolve_tac [allI, ballI, impI])))))) x;
berghofe@7717
   765
berghofe@7717
   766
in
berghofe@7717
   767
berghofe@7717
   768
val rulify_prems_attrib_setup =
berghofe@7717
   769
 [Attrib.add_attributes
berghofe@7717
   770
  [("rulify_prems", (gen_rulify_prems, gen_rulify_prems), "put theorem into standard rule form")]];
berghofe@7717
   771
berghofe@7717
   772
end;