src/HOL/Set.ML
author paulson
Tue, 16 Jul 1996 15:49:46 +0200
changeset 1868 836950047d85
parent 1841 8e5e2fef6d26
child 1882 67f49e8c4355
permissions -rw-r--r--
Put in minimal simpset to avoid excessive simplification, just as in revision 1.9 of HOL/indrule.ML
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
     1
(*  Title:      HOL/set
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
For set.thy.  Set theory for higher-order logic.  A set is simply a predicate.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
open Set;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    11
section "Relating predicates and sets";
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    12
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    13
val [prem] = goal Set.thy "P(a) ==> a : {x.P(x)}";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
by (rtac (mem_Collect_eq RS ssubst) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
by (rtac prem 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
qed "CollectI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
val prems = goal Set.thy "[| a : {x.P(x)} |] ==> P(a)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
by (resolve_tac (prems RL [mem_Collect_eq  RS subst]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
qed "CollectD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
val [prem] = goal Set.thy "[| !!x. (x:A) = (x:B) |] ==> A = B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
by (rtac (prem RS ext RS arg_cong RS box_equals) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
by (rtac Collect_mem_eq 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
by (rtac Collect_mem_eq 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
qed "set_ext";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
val [prem] = goal Set.thy "[| !!x. P(x)=Q(x) |] ==> {x. P(x)} = {x. Q(x)}";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
by (rtac (prem RS ext RS arg_cong) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
qed "Collect_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
val CollectE = make_elim CollectD;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    34
section "Bounded quantifiers";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
val prems = goalw Set.thy [Ball_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
    "[| !!x. x:A ==> P(x) |] ==> ! x:A. P(x)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
by (REPEAT (ares_tac (prems @ [allI,impI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
qed "ballI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
val [major,minor] = goalw Set.thy [Ball_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
    "[| ! x:A. P(x);  x:A |] ==> P(x)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
by (rtac (minor RS (major RS spec RS mp)) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
qed "bspec";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
val major::prems = goalw Set.thy [Ball_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
    "[| ! x:A. P(x);  P(x) ==> Q;  x~:A ==> Q |] ==> Q";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
by (rtac (major RS spec RS impCE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
qed "ballE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
(*Takes assumptions ! x:A.P(x) and a:A; creates assumption P(a)*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
fun ball_tac i = etac ballE i THEN contr_tac (i+1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
val prems = goalw Set.thy [Bex_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
    "[| P(x);  x:A |] ==> ? x:A. P(x)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
by (REPEAT (ares_tac (prems @ [exI,conjI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
qed "bexI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
qed_goal "bexCI" Set.thy 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
   "[| ! x:A. ~P(x) ==> P(a);  a:A |] ==> ? x:A.P(x)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
 (fn prems=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
  [ (rtac classical 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
    (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1))  ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
val major::prems = goalw Set.thy [Bex_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
    "[| ? x:A. P(x);  !!x. [| x:A; P(x) |] ==> Q  |] ==> Q";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
by (rtac (major RS exE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
qed "bexE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
(*Trival rewrite rule;   (! x:A.P)=P holds only if A is nonempty!*)
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1552
diff changeset
    73
goal Set.thy "(! x:A. True) = True";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
by (REPEAT (ares_tac [TrueI,ballI,iffI] 1));
1816
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
    75
qed "ball_True";
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
    76
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
    77
Addsimps [ball_True];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
(** Congruence rules **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
    "[| A=B;  !!x. x:B ==> P(x) = Q(x) |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
\    (! x:A. P(x)) = (! x:B. Q(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
by (resolve_tac (prems RL [ssubst]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
by (REPEAT (ares_tac [ballI,iffI] 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
     ORELSE eresolve_tac ([make_elim bspec, mp] @ (prems RL [iffE])) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
qed "ball_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
    "[| A=B;  !!x. x:B ==> P(x) = Q(x) |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
\    (? x:A. P(x)) = (? x:B. Q(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    92
by (resolve_tac (prems RL [ssubst]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
by (REPEAT (etac bexE 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
     ORELSE ares_tac ([bexI,iffI] @ (prems RL [iffD1,iffD2])) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
qed "bex_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    97
section "Subsets";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
val prems = goalw Set.thy [subset_def] "(!!x.x:A ==> x:B) ==> A <= B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
by (REPEAT (ares_tac (prems @ [ballI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
qed "subsetI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
(*Rule in Modus Ponens style*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
val major::prems = goalw Set.thy [subset_def] "[| A <= B;  c:A |] ==> c:B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
by (rtac (major RS bspec) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
qed "subsetD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
(*The same, with reversed premises for use with etac -- cf rev_mp*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
qed_goal "rev_subsetD" Set.thy "[| c:A;  A <= B |] ==> c:B"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
 (fn prems=>  [ (REPEAT (resolve_tac (prems@[subsetD]) 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
1841
8e5e2fef6d26 Added contra_subsetD and rev_contra_subsetD
paulson
parents: 1816
diff changeset
   113
qed_goal "contra_subsetD" Set.thy "!!c. [| A <= B; c ~: B |] ==> c ~: A"
8e5e2fef6d26 Added contra_subsetD and rev_contra_subsetD
paulson
parents: 1816
diff changeset
   114
 (fn prems=>  [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]);
8e5e2fef6d26 Added contra_subsetD and rev_contra_subsetD
paulson
parents: 1816
diff changeset
   115
8e5e2fef6d26 Added contra_subsetD and rev_contra_subsetD
paulson
parents: 1816
diff changeset
   116
qed_goal "rev_contra_subsetD" Set.thy "!!c. [| c ~: B;  A <= B |] ==> c ~: A"
8e5e2fef6d26 Added contra_subsetD and rev_contra_subsetD
paulson
parents: 1816
diff changeset
   117
 (fn prems=>  [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]);
8e5e2fef6d26 Added contra_subsetD and rev_contra_subsetD
paulson
parents: 1816
diff changeset
   118
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
(*Classical elimination rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
val major::prems = goalw Set.thy [subset_def] 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
    "[| A <= B;  c~:A ==> P;  c:B ==> P |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
by (rtac (major RS ballE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   124
qed "subsetCE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
(*Takes assumptions A<=B; c:A and creates the assumption c:B *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
fun set_mp_tac i = etac subsetCE i  THEN  mp_tac i;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
qed_goal "subset_refl" Set.thy "A <= (A::'a set)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
 (fn _=> [ (REPEAT (ares_tac [subsetI] 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   131
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
val prems = goal Set.thy "[| A<=B;  B<=C |] ==> A<=(C::'a set)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
by (cut_facts_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
by (REPEAT (ares_tac [subsetI] 1 ORELSE set_mp_tac 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
qed "subset_trans";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   138
section "Equality";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
(*Anti-symmetry of the subset relation*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   141
val prems = goal Set.thy "[| A <= B;  B <= A |] ==> A = (B::'a set)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   142
by (rtac (iffI RS set_ext) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   143
by (REPEAT (ares_tac (prems RL [subsetD]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   144
qed "subset_antisym";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   145
val equalityI = subset_antisym;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   146
1762
6e481897a811 equalityI is now added to default claset
berghofe
parents: 1760
diff changeset
   147
AddSIs [equalityI];
6e481897a811 equalityI is now added to default claset
berghofe
parents: 1760
diff changeset
   148
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   149
(* Equality rules from ZF set theory -- are they appropriate here? *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   150
val prems = goal Set.thy "A = B ==> A<=(B::'a set)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   151
by (resolve_tac (prems RL [subst]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   152
by (rtac subset_refl 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   153
qed "equalityD1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   154
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   155
val prems = goal Set.thy "A = B ==> B<=(A::'a set)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   156
by (resolve_tac (prems RL [subst]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   157
by (rtac subset_refl 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   158
qed "equalityD2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   159
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   160
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   161
    "[| A = B;  [| A<=B; B<=(A::'a set) |] ==> P |]  ==>  P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   162
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   163
by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   164
qed "equalityE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   165
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   166
val major::prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   167
    "[| A = B;  [| c:A; c:B |] ==> P;  [| c~:A; c~:B |] ==> P |]  ==>  P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   168
by (rtac (major RS equalityE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   169
by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   170
qed "equalityCE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   171
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   172
(*Lemma for creating induction formulae -- for "pattern matching" on p
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   173
  To make the induction hypotheses usable, apply "spec" or "bspec" to
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   174
  put universal quantifiers over the free variables in p. *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   175
val prems = goal Set.thy 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   176
    "[| p:A;  !!z. z:A ==> p=z --> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
by (rtac mp 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
by (REPEAT (resolve_tac (refl::prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   179
qed "setup_induction";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   180
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   181
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   182
section "Set complement -- Compl";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   183
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   184
val prems = goalw Set.thy [Compl_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   185
    "[| c:A ==> False |] ==> c : Compl(A)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   186
by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   187
qed "ComplI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   188
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   189
(*This form, with negated conclusion, works well with the Classical prover.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   190
  Negated assumptions behave like formulae on the right side of the notional
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   191
  turnstile...*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
val major::prems = goalw Set.thy [Compl_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   193
    "[| c : Compl(A) |] ==> c~:A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   194
by (rtac (major RS CollectD) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
qed "ComplD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   196
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   197
val ComplE = make_elim ComplD;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   199
qed_goal "Compl_iff" Set.thy "(c : Compl(A)) = (c~:A)"
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1640
diff changeset
   200
 (fn _ => [ (fast_tac (!claset addSIs [ComplI] addSEs [ComplE]) 1) ]);
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   201
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   202
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   203
section "Binary union -- Un";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   204
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   205
val prems = goalw Set.thy [Un_def] "c:A ==> c : A Un B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   206
by (REPEAT (resolve_tac (prems @ [CollectI,disjI1]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   207
qed "UnI1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   208
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   209
val prems = goalw Set.thy [Un_def] "c:B ==> c : A Un B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   210
by (REPEAT (resolve_tac (prems @ [CollectI,disjI2]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   211
qed "UnI2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   212
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   213
(*Classical introduction rule: no commitment to A vs B*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   214
qed_goal "UnCI" Set.thy "(c~:B ==> c:A) ==> c : A Un B"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   215
 (fn prems=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   216
  [ (rtac classical 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   217
    (REPEAT (ares_tac (prems@[UnI1,notI]) 1)),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   218
    (REPEAT (ares_tac (prems@[UnI2,notE]) 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   219
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   220
val major::prems = goalw Set.thy [Un_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   221
    "[| c : A Un B;  c:A ==> P;  c:B ==> P |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   222
by (rtac (major RS CollectD RS disjE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   223
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   224
qed "UnE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   225
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   226
qed_goal "Un_iff" Set.thy "(c : A Un B) = (c:A | c:B)"
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1640
diff changeset
   227
 (fn _ => [ (fast_tac (!claset addSIs [UnCI] addSEs [UnE]) 1) ]);
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   228
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   229
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   230
section "Binary intersection -- Int";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   231
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   232
val prems = goalw Set.thy [Int_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   233
    "[| c:A;  c:B |] ==> c : A Int B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   234
by (REPEAT (resolve_tac (prems @ [CollectI,conjI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   235
qed "IntI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   236
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   237
val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   238
by (rtac (major RS CollectD RS conjunct1) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   239
qed "IntD1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   240
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   241
val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   242
by (rtac (major RS CollectD RS conjunct2) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   243
qed "IntD2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   244
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   245
val [major,minor] = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   246
    "[| c : A Int B;  [| c:A; c:B |] ==> P |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   247
by (rtac minor 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   248
by (rtac (major RS IntD1) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   249
by (rtac (major RS IntD2) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   250
qed "IntE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   251
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   252
qed_goal "Int_iff" Set.thy "(c : A Int B) = (c:A & c:B)"
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1640
diff changeset
   253
 (fn _ => [ (fast_tac (!claset addSIs [IntI] addSEs [IntE]) 1) ]);
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   254
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   255
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   256
section "Set difference";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   257
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   258
qed_goalw "DiffI" Set.thy [set_diff_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   259
    "[| c : A;  c ~: B |] ==> c : A - B"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   260
 (fn prems=> [ (REPEAT (resolve_tac (prems @ [CollectI,conjI]) 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   261
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   262
qed_goalw "DiffD1" Set.thy [set_diff_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   263
    "c : A - B ==> c : A"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   264
 (fn [major]=> [ (rtac (major RS CollectD RS conjunct1) 1) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   265
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   266
qed_goalw "DiffD2" Set.thy [set_diff_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   267
    "[| c : A - B;  c : B |] ==> P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   268
 (fn [major,minor]=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   269
     [rtac (minor RS (major RS CollectD RS conjunct2 RS notE)) 1]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   270
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   271
qed_goal "DiffE" Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   272
    "[| c : A - B;  [| c:A; c~:B |] ==> P |] ==> P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   273
 (fn prems=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   274
  [ (resolve_tac prems 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   275
    (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   276
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   277
qed_goal "Diff_iff" Set.thy "(c : A-B) = (c:A & c~:B)"
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1640
diff changeset
   278
 (fn _ => [ (fast_tac (!claset addSIs [DiffI] addSEs [DiffE]) 1) ]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   279
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   280
section "The empty set -- {}";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   281
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   282
qed_goalw "emptyE" Set.thy [empty_def] "a:{} ==> P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   283
 (fn [prem] => [rtac (prem RS CollectD RS FalseE) 1]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   284
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   285
qed_goal "empty_subsetI" Set.thy "{} <= A"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   286
 (fn _ => [ (REPEAT (ares_tac [equalityI,subsetI,emptyE] 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   287
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   288
qed_goal "equals0I" Set.thy "[| !!y. y:A ==> False |] ==> A={}"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   289
 (fn prems=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   290
  [ (REPEAT (ares_tac (prems@[empty_subsetI,subsetI,equalityI]) 1 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   291
      ORELSE eresolve_tac (prems RL [FalseE]) 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   292
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   293
qed_goal "equals0D" Set.thy "[| A={};  a:A |] ==> P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   294
 (fn [major,minor]=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   295
  [ (rtac (minor RS (major RS equalityD1 RS subsetD RS emptyE)) 1) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   296
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   297
qed_goal "empty_iff" Set.thy "(c : {}) = False"
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1640
diff changeset
   298
 (fn _ => [ (fast_tac (!claset addSEs [emptyE]) 1) ]);
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   299
1816
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
   300
goal Set.thy "Ball {} P = True";
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
   301
by (simp_tac (HOL_ss addsimps [mem_Collect_eq, Ball_def, empty_def]) 1);
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
   302
qed "ball_empty";
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
   303
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
   304
goal Set.thy "Bex {} P = False";
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
   305
by (simp_tac (HOL_ss addsimps [mem_Collect_eq, Bex_def, empty_def]) 1);
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
   306
qed "bex_empty";
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
   307
Addsimps [ball_empty, bex_empty];
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
   308
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   309
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   310
section "Augmenting a set -- insert";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   311
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   312
qed_goalw "insertI1" Set.thy [insert_def] "a : insert a B"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   313
 (fn _ => [rtac (CollectI RS UnI1) 1, rtac refl 1]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   314
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   315
qed_goalw "insertI2" Set.thy [insert_def] "a : B ==> a : insert b B"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   316
 (fn [prem]=> [ (rtac (prem RS UnI2) 1) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   317
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   318
qed_goalw "insertE" Set.thy [insert_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   319
    "[| a : insert b A;  a=b ==> P;  a:A ==> P |] ==> P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   320
 (fn major::prems=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   321
  [ (rtac (major RS UnE) 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   322
    (REPEAT (eresolve_tac (prems @ [CollectE]) 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   323
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   324
qed_goal "insert_iff" Set.thy "a : insert b A = (a=b | a:A)"
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1640
diff changeset
   325
 (fn _ => [fast_tac (!claset addIs [insertI1,insertI2] addSEs [insertE]) 1]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   326
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   327
(*Classical introduction rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   328
qed_goal "insertCI" Set.thy "(a~:B ==> a=b) ==> a: insert b B"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   329
 (fn [prem]=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   330
  [ (rtac (disjCI RS (insert_iff RS iffD2)) 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   331
    (etac prem 1) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   332
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   333
section "Singletons, using insert";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   334
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   335
qed_goal "singletonI" Set.thy "a : {a}"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   336
 (fn _=> [ (rtac insertI1 1) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   337
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   338
goalw Set.thy [insert_def] "!!a. b : {a} ==> b=a";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1640
diff changeset
   339
by (fast_tac (!claset addSEs [emptyE,CollectE,UnE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   340
qed "singletonD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   341
1776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   342
bind_thm ("singletonE", make_elim singletonD);
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   343
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   344
qed_goal "singleton_iff" thy "(b : {a}) = (b=a)" (fn _ => [
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   345
	rtac iffI 1,
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   346
	etac singletonD 1,
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   347
	hyp_subst_tac 1,
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   348
	rtac singletonI 1]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   349
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   350
val [major] = goal Set.thy "{a}={b} ==> a=b";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   351
by (rtac (major RS equalityD1 RS subsetD RS singletonD) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   352
by (rtac singletonI 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   353
qed "singleton_inject";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   354
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   355
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   356
section "The universal set -- UNIV";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   357
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   358
qed_goal "subset_UNIV" Set.thy "A <= UNIV"
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   359
  (fn _ => [rtac subsetI 1, rtac ComplI 1, etac emptyE 1]);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   360
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   361
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   362
section "Unions of families -- UNION x:A. B(x) is Union(B``A)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   363
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   364
(*The order of the premises presupposes that A is rigid; b may be flexible*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   365
val prems = goalw Set.thy [UNION_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   366
    "[| a:A;  b: B(a) |] ==> b: (UN x:A. B(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   367
by (REPEAT (resolve_tac (prems @ [bexI,CollectI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   368
qed "UN_I";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   369
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   370
val major::prems = goalw Set.thy [UNION_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   371
    "[| b : (UN x:A. B(x));  !!x.[| x:A;  b: B(x) |] ==> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   372
by (rtac (major RS CollectD RS bexE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   373
by (REPEAT (ares_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   374
qed "UN_E";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   375
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   376
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   377
    "[| A=B;  !!x. x:B ==> C(x) = D(x) |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   378
\    (UN x:A. C(x)) = (UN x:B. D(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   379
by (REPEAT (etac UN_E 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   380
     ORELSE ares_tac ([UN_I,equalityI,subsetI] @ 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
   381
                      (prems RL [equalityD1,equalityD2] RL [subsetD])) 1));
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   382
qed "UN_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   383
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   384
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   385
section "Intersections of families -- INTER x:A. B(x) is Inter(B``A)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   386
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   387
val prems = goalw Set.thy [INTER_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   388
    "(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   389
by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   390
qed "INT_I";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   391
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   392
val major::prems = goalw Set.thy [INTER_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   393
    "[| b : (INT x:A. B(x));  a:A |] ==> b: B(a)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   394
by (rtac (major RS CollectD RS bspec) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   395
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   396
qed "INT_D";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   397
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   398
(*"Classical" elimination -- by the Excluded Middle on a:A *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   399
val major::prems = goalw Set.thy [INTER_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   400
    "[| b : (INT x:A. B(x));  b: B(a) ==> R;  a~:A ==> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   401
by (rtac (major RS CollectD RS ballE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   402
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   403
qed "INT_E";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   404
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   405
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   406
    "[| A=B;  !!x. x:B ==> C(x) = D(x) |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   407
\    (INT x:A. C(x)) = (INT x:B. D(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   408
by (REPEAT_FIRST (resolve_tac [INT_I,equalityI,subsetI]));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   409
by (REPEAT (dtac INT_D 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   410
     ORELSE ares_tac (prems RL [equalityD1,equalityD2] RL [subsetD]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   411
qed "INT_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   412
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   413
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   414
section "Unions over a type; UNION1(B) = Union(range(B))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   415
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   416
(*The order of the premises presupposes that A is rigid; b may be flexible*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   417
val prems = goalw Set.thy [UNION1_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   418
    "b: B(x) ==> b: (UN x. B(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   419
by (REPEAT (resolve_tac (prems @ [TrueI, CollectI RS UN_I]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   420
qed "UN1_I";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   421
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   422
val major::prems = goalw Set.thy [UNION1_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   423
    "[| b : (UN x. B(x));  !!x. b: B(x) ==> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   424
by (rtac (major RS UN_E) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   425
by (REPEAT (ares_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   426
qed "UN1_E";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   427
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   428
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   429
section "Intersections over a type; INTER1(B) = Inter(range(B))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   430
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   431
val prems = goalw Set.thy [INTER1_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   432
    "(!!x. b: B(x)) ==> b : (INT x. B(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   433
by (REPEAT (ares_tac (INT_I::prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   434
qed "INT1_I";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   435
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   436
val [major] = goalw Set.thy [INTER1_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   437
    "b : (INT x. B(x)) ==> b: B(a)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   438
by (rtac (TrueI RS (CollectI RS (major RS INT_D))) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   439
qed "INT1_D";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   440
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   441
section "Union";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   442
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   443
(*The order of the premises presupposes that C is rigid; A may be flexible*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   444
val prems = goalw Set.thy [Union_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   445
    "[| X:C;  A:X |] ==> A : Union(C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   446
by (REPEAT (resolve_tac (prems @ [UN_I]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   447
qed "UnionI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   448
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   449
val major::prems = goalw Set.thy [Union_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   450
    "[| A : Union(C);  !!X.[| A:X;  X:C |] ==> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   451
by (rtac (major RS UN_E) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   452
by (REPEAT (ares_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   453
qed "UnionE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   454
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   455
section "Inter";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   456
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   457
val prems = goalw Set.thy [Inter_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   458
    "[| !!X. X:C ==> A:X |] ==> A : Inter(C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   459
by (REPEAT (ares_tac ([INT_I] @ prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   460
qed "InterI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   461
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   462
(*A "destruct" rule -- every X in C contains A as an element, but
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   463
  A:X can hold when X:C does not!  This rule is analogous to "spec". *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   464
val major::prems = goalw Set.thy [Inter_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   465
    "[| A : Inter(C);  X:C |] ==> A:X";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   466
by (rtac (major RS INT_D) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   467
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   468
qed "InterD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   469
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   470
(*"Classical" elimination rule -- does not require proving X:C *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   471
val major::prems = goalw Set.thy [Inter_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   472
    "[| A : Inter(C);  A:X ==> R;  X~:C ==> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   473
by (rtac (major RS INT_E) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   474
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   475
qed "InterE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   476
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   477
section "The Powerset operator -- Pow";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   478
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   479
qed_goalw "PowI" Set.thy [Pow_def] "!!A B. A <= B ==> A : Pow(B)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   480
 (fn _ => [ (etac CollectI 1) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   481
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   482
qed_goalw "PowD" Set.thy [Pow_def] "!!A B. A : Pow(B)  ==>  A<=B"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   483
 (fn _=> [ (etac CollectD 1) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   484
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   485
val Pow_bottom = empty_subsetI RS PowI;        (* {}: Pow(B) *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   486
val Pow_top = subset_refl RS PowI;             (* A : Pow(A) *)
1776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   487
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   488
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   489
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   490
(*** Set reasoning tools ***)
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   491
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   492
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   493
val mem_simps = [ Un_iff, Int_iff, Compl_iff, Diff_iff, singleton_iff,
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   494
		  mem_Collect_eq];
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   495
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   496
val mksimps_pairs = ("Ball",[bspec]) :: mksimps_pairs;
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   497
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   498
simpset := !simpset addsimps mem_simps
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   499
                    addcongs [ball_cong,bex_cong]
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   500
                    setmksimps (mksimps mksimps_pairs);