src/HOL/Set.ML
author paulson
Sat, 01 Nov 1997 12:59:06 +0100
changeset 4059 59c1422c9da5
parent 3960 7a38fae985f9
child 4089 96fba19bcbe2
permissions -rw-r--r--
New Blast_tac (and minor tidying...)
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
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1937
diff changeset
     6
Set theory for higher-order logic.  A set is simply a predicate.
923
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
3469
61d927bd57ec Now Collect_mem_eq is a default simprule (how could it have ever been omitted?
paulson
parents: 3420
diff changeset
    13
Addsimps [Collect_mem_eq];
61d927bd57ec Now Collect_mem_eq is a default simprule (how could it have ever been omitted?
paulson
parents: 3420
diff changeset
    14
AddIffs  [mem_Collect_eq];
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    15
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3718
diff changeset
    16
goal Set.thy "!!a. P(a) ==> a : {x. P(x)}";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    17
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
qed "CollectI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3718
diff changeset
    20
val prems = goal Set.thy "!!a. a : {x. P(x)} ==> P(a)";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    21
by (Asm_full_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
qed "CollectD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
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
    25
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
    26
by (rtac Collect_mem_eq 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
by (rtac Collect_mem_eq 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
qed "set_ext";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
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
    31
by (rtac (prem RS ext RS arg_cong) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
qed "Collect_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
val CollectE = make_elim CollectD;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    36
AddSIs [CollectI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    37
AddSEs [CollectE];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    38
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    39
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    40
section "Bounded quantifiers";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
val prems = goalw Set.thy [Ball_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
    "[| !!x. x:A ==> P(x) |] ==> ! x:A. P(x)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
by (REPEAT (ares_tac (prems @ [allI,impI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
qed "ballI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
val [major,minor] = goalw Set.thy [Ball_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
    "[| ! x:A. P(x);  x:A |] ==> P(x)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
by (rtac (minor RS (major RS spec RS mp)) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
qed "bspec";
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
val major::prems = goalw Set.thy [Ball_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
    "[| ! x:A. P(x);  P(x) ==> Q;  x~:A ==> Q |] ==> Q";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
by (rtac (major RS spec RS impCE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
qed "ballE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
(*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
    59
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
    60
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    61
AddSIs [ballI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    62
AddEs  [ballE];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    63
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
val prems = goalw Set.thy [Bex_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
    "[| P(x);  x:A |] ==> ? x:A. P(x)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
by (REPEAT (ares_tac (prems @ [exI,conjI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
qed "bexI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
qed_goal "bexCI" Set.thy 
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3718
diff changeset
    70
   "[| ! x:A. ~P(x) ==> P(a);  a:A |] ==> ? x:A. P(x)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
 (fn prems=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
  [ (rtac classical 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
    (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1))  ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
val major::prems = goalw Set.thy [Bex_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
    "[| ? x:A. P(x);  !!x. [| x:A; P(x) |] ==> Q  |] ==> Q";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
by (rtac (major RS exE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
qed "bexE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    81
AddIs  [bexI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    82
AddSEs [bexE];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    83
3420
02dc9c5b035f New miniscoping rules ball_triv and bex_triv
paulson
parents: 3222
diff changeset
    84
(*Trival rewrite rule*)
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3718
diff changeset
    85
goal Set.thy "(! x:A. P) = ((? x. x:A) --> P)";
3420
02dc9c5b035f New miniscoping rules ball_triv and bex_triv
paulson
parents: 3222
diff changeset
    86
by (simp_tac (!simpset addsimps [Ball_def]) 1);
02dc9c5b035f New miniscoping rules ball_triv and bex_triv
paulson
parents: 3222
diff changeset
    87
qed "ball_triv";
1816
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
    88
1882
67f49e8c4355 Proved bex_False
paulson
parents: 1841
diff changeset
    89
(*Dual form for existentials*)
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3718
diff changeset
    90
goal Set.thy "(? x:A. P) = ((? x. x:A) & P)";
3420
02dc9c5b035f New miniscoping rules ball_triv and bex_triv
paulson
parents: 3222
diff changeset
    91
by (simp_tac (!simpset addsimps [Bex_def]) 1);
02dc9c5b035f New miniscoping rules ball_triv and bex_triv
paulson
parents: 3222
diff changeset
    92
qed "bex_triv";
1882
67f49e8c4355 Proved bex_False
paulson
parents: 1841
diff changeset
    93
3420
02dc9c5b035f New miniscoping rules ball_triv and bex_triv
paulson
parents: 3222
diff changeset
    94
Addsimps [ball_triv, bex_triv];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
(** Congruence rules **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
    "[| A=B;  !!x. x:B ==> P(x) = Q(x) |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
\    (! x:A. P(x)) = (! x:B. Q(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
by (resolve_tac (prems RL [ssubst]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
by (REPEAT (ares_tac [ballI,iffI] 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
     ORELSE eresolve_tac ([make_elim bspec, mp] @ (prems RL [iffE])) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
qed "ball_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
    "[| A=B;  !!x. x:B ==> P(x) = Q(x) |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
\    (? x:A. P(x)) = (? x:B. Q(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
by (resolve_tac (prems RL [ssubst]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
by (REPEAT (etac bexE 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
     ORELSE ares_tac ([bexI,iffI] @ (prems RL [iffD1,iffD2])) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
qed "bex_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   114
section "Subsets";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3718
diff changeset
   116
val prems = goalw Set.thy [subset_def] "(!!x. x:A ==> x:B) ==> A <= B";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
by (REPEAT (ares_tac (prems @ [ballI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   118
qed "subsetI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   120
Blast.overload ("op <=", domain_type); (*The <= relation is overloaded*)
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   121
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   122
(*While (:) is not, its type must be kept
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   123
  for overloading of = to work.*)
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   124
Blast.overload ("op :", domain_type);
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   125
seq (fn a => Blast.overload (a, HOLogic.dest_setT o domain_type))
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   126
    ["Ball", "Bex"];
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   127
(*need UNION, INTER also?*)
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   128
2881
62ecde1015ae Declares overloading for set-theoretic constants
paulson
parents: 2858
diff changeset
   129
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
(*Rule in Modus Ponens style*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   131
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
   132
by (rtac (major RS bspec) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
qed "subsetD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
(*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
   137
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
   138
 (fn prems=>  [ (REPEAT (resolve_tac (prems@[subsetD]) 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
1920
df683ce7aad8 Added impOfSubs
paulson
parents: 1882
diff changeset
   140
(*Converts A<=B to x:A ==> x:B*)
df683ce7aad8 Added impOfSubs
paulson
parents: 1882
diff changeset
   141
fun impOfSubs th = th RSN (2, rev_subsetD);
df683ce7aad8 Added impOfSubs
paulson
parents: 1882
diff changeset
   142
1841
8e5e2fef6d26 Added contra_subsetD and rev_contra_subsetD
paulson
parents: 1816
diff changeset
   143
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
   144
 (fn prems=>  [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]);
8e5e2fef6d26 Added contra_subsetD and rev_contra_subsetD
paulson
parents: 1816
diff changeset
   145
8e5e2fef6d26 Added contra_subsetD and rev_contra_subsetD
paulson
parents: 1816
diff changeset
   146
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
   147
 (fn prems=>  [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]);
8e5e2fef6d26 Added contra_subsetD and rev_contra_subsetD
paulson
parents: 1816
diff changeset
   148
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   149
(*Classical elimination rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   150
val major::prems = goalw Set.thy [subset_def] 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   151
    "[| A <= B;  c~:A ==> P;  c:B ==> P |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   152
by (rtac (major RS ballE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   153
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   154
qed "subsetCE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   155
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   156
(*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
   157
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
   158
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   159
AddSIs [subsetI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   160
AddEs  [subsetD, subsetCE];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   161
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   162
qed_goal "subset_refl" Set.thy "A <= (A::'a set)"
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   163
 (fn _=> [Fast_tac 1]);		(*Blast_tac would try order_refl and fail*)
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   164
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   165
val prems = goal Set.thy "!!B. [| A<=B;  B<=C |] ==> A<=(C::'a set)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   166
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   167
qed "subset_trans";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   168
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   169
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   170
section "Equality";
923
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
(*Anti-symmetry of the subset relation*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   173
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
   174
by (rtac (iffI RS set_ext) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   175
by (REPEAT (ares_tac (prems RL [subsetD]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   176
qed "subset_antisym";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
val equalityI = subset_antisym;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
1762
6e481897a811 equalityI is now added to default claset
berghofe
parents: 1760
diff changeset
   179
AddSIs [equalityI];
6e481897a811 equalityI is now added to default claset
berghofe
parents: 1760
diff changeset
   180
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   181
(* Equality rules from ZF set theory -- are they appropriate here? *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   182
val prems = goal Set.thy "A = B ==> A<=(B::'a set)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   183
by (resolve_tac (prems RL [subst]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   184
by (rtac subset_refl 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   185
qed "equalityD1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   186
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   187
val prems = goal Set.thy "A = B ==> B<=(A::'a set)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   188
by (resolve_tac (prems RL [subst]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   189
by (rtac subset_refl 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   190
qed "equalityD2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   191
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   193
    "[| A = B;  [| A<=B; B<=(A::'a set) |] ==> P |]  ==>  P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   194
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   196
qed "equalityE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   197
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
val major::prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   199
    "[| 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
   200
by (rtac (major RS equalityE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   201
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
   202
qed "equalityCE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   203
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   204
(*Lemma for creating induction formulae -- for "pattern matching" on p
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   205
  To make the induction hypotheses usable, apply "spec" or "bspec" to
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   206
  put universal quantifiers over the free variables in p. *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   207
val prems = goal Set.thy 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   208
    "[| p:A;  !!z. z:A ==> p=z --> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   209
by (rtac mp 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   210
by (REPEAT (resolve_tac (refl::prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   211
qed "setup_induction";
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
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   214
section "The empty set -- {}";
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   215
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   216
qed_goalw "empty_iff" Set.thy [empty_def] "(c : {}) = False"
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   217
 (fn _ => [ (Blast_tac 1) ]);
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   218
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   219
Addsimps [empty_iff];
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   220
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   221
qed_goal "emptyE" Set.thy "!!a. a:{} ==> P"
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   222
 (fn _ => [Full_simp_tac 1]);
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   223
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   224
AddSEs [emptyE];
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   225
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   226
qed_goal "empty_subsetI" Set.thy "{} <= A"
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   227
 (fn _ => [ (Blast_tac 1) ]);
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   228
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   229
qed_goal "equals0I" Set.thy "[| !!y. y:A ==> False |] ==> A={}"
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   230
 (fn [prem]=>
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2912
diff changeset
   231
  [ (blast_tac (!claset addIs [prem RS FalseE]) 1) ]);
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   232
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   233
qed_goal "equals0D" Set.thy "!!a. [| A={};  a:A |] ==> P"
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   234
 (fn _ => [ (Blast_tac 1) ]);
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   235
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   236
goal Set.thy "Ball {} P = True";
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   237
by (simp_tac (HOL_ss addsimps [mem_Collect_eq, Ball_def, empty_def]) 1);
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   238
qed "ball_empty";
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   239
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   240
goal Set.thy "Bex {} P = False";
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   241
by (simp_tac (HOL_ss addsimps [mem_Collect_eq, Bex_def, empty_def]) 1);
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   242
qed "bex_empty";
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   243
Addsimps [ball_empty, bex_empty];
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   244
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   245
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   246
section "The Powerset operator -- Pow";
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   247
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   248
qed_goalw "Pow_iff" Set.thy [Pow_def] "(A : Pow(B)) = (A <= B)"
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   249
 (fn _ => [ (Asm_simp_tac 1) ]);
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   250
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   251
AddIffs [Pow_iff]; 
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   252
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   253
qed_goalw "PowI" Set.thy [Pow_def] "!!A B. A <= B ==> A : Pow(B)"
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   254
 (fn _ => [ (etac CollectI 1) ]);
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   255
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   256
qed_goalw "PowD" Set.thy [Pow_def] "!!A B. A : Pow(B)  ==>  A<=B"
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   257
 (fn _=> [ (etac CollectD 1) ]);
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   258
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   259
val Pow_bottom = empty_subsetI RS PowI;        (* {}: Pow(B) *)
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   260
val Pow_top = subset_refl RS PowI;             (* A : Pow(A) *)
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   261
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   262
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   263
section "Set complement -- Compl";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   264
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   265
qed_goalw "Compl_iff" Set.thy [Compl_def] "(c : Compl(A)) = (c~:A)"
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   266
 (fn _ => [ (Blast_tac 1) ]);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   267
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   268
Addsimps [Compl_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   269
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   270
val prems = goalw Set.thy [Compl_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   271
    "[| c:A ==> False |] ==> c : Compl(A)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   272
by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   273
qed "ComplI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   274
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   275
(*This form, with negated conclusion, works well with the Classical prover.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   276
  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
   277
  turnstile...*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   278
val major::prems = goalw Set.thy [Compl_def]
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   279
    "c : Compl(A) ==> c~:A";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   280
by (rtac (major RS CollectD) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   281
qed "ComplD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   282
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   283
val ComplE = make_elim ComplD;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   284
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   285
AddSIs [ComplI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   286
AddSEs [ComplE];
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   287
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   288
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   289
section "Binary union -- Un";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   290
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   291
qed_goalw "Un_iff" Set.thy [Un_def] "(c : A Un B) = (c:A | c:B)"
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   292
 (fn _ => [ Blast_tac 1 ]);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   293
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   294
Addsimps [Un_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   295
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   296
goal Set.thy "!!c. c:A ==> c : A Un B";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   297
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   298
qed "UnI1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   299
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   300
goal Set.thy "!!c. c:B ==> c : A Un B";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   301
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   302
qed "UnI2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   303
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   304
(*Classical introduction rule: no commitment to A vs B*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   305
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
   306
 (fn prems=>
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   307
  [ (Simp_tac 1),
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   308
    (REPEAT (ares_tac (prems@[disjCI]) 1)) ]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   309
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   310
val major::prems = goalw Set.thy [Un_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   311
    "[| c : A Un B;  c:A ==> P;  c:B ==> P |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   312
by (rtac (major RS CollectD RS disjE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   313
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   314
qed "UnE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   315
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   316
AddSIs [UnCI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   317
AddSEs [UnE];
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   318
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   319
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   320
section "Binary intersection -- Int";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   321
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   322
qed_goalw "Int_iff" Set.thy [Int_def] "(c : A Int B) = (c:A & c:B)"
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   323
 (fn _ => [ (Blast_tac 1) ]);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   324
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   325
Addsimps [Int_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   326
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   327
goal Set.thy "!!c. [| c:A;  c:B |] ==> c : A Int B";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   328
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   329
qed "IntI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   330
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   331
goal Set.thy "!!c. c : A Int B ==> c:A";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   332
by (Asm_full_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   333
qed "IntD1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   334
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   335
goal Set.thy "!!c. c : A Int B ==> c:B";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   336
by (Asm_full_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   337
qed "IntD2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   338
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   339
val [major,minor] = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   340
    "[| c : A Int B;  [| c:A; c:B |] ==> P |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   341
by (rtac minor 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   342
by (rtac (major RS IntD1) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   343
by (rtac (major RS IntD2) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   344
qed "IntE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   345
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   346
AddSIs [IntI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   347
AddSEs [IntE];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   348
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   349
section "Set difference";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   350
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   351
qed_goalw "Diff_iff" Set.thy [set_diff_def] "(c : A-B) = (c:A & c~:B)"
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   352
 (fn _ => [ (Blast_tac 1) ]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   353
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   354
Addsimps [Diff_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   355
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   356
qed_goal "DiffI" Set.thy "!!c. [| c : A;  c ~: B |] ==> c : A - B"
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   357
 (fn _=> [ Asm_simp_tac 1 ]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   358
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   359
qed_goal "DiffD1" Set.thy "!!c. c : A - B ==> c : A"
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   360
 (fn _=> [ (Asm_full_simp_tac 1) ]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   361
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   362
qed_goal "DiffD2" Set.thy "!!c. [| c : A - B;  c : B |] ==> P"
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   363
 (fn _=> [ (Asm_full_simp_tac 1) ]);
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   364
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   365
qed_goal "DiffE" Set.thy "[| c : A - B;  [| c:A; c~:B |] ==> P |] ==> P"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   366
 (fn prems=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   367
  [ (resolve_tac prems 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   368
    (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   369
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   370
AddSIs [DiffI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   371
AddSEs [DiffE];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   372
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   373
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   374
section "Augmenting a set -- insert";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   375
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   376
qed_goalw "insert_iff" Set.thy [insert_def] "a : insert b A = (a=b | a:A)"
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   377
 (fn _ => [Blast_tac 1]);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   378
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   379
Addsimps [insert_iff];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   380
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   381
qed_goal "insertI1" Set.thy "a : insert a B"
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   382
 (fn _ => [Simp_tac 1]);
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   383
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   384
qed_goal "insertI2" Set.thy "!!a. a : B ==> a : insert b B"
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   385
 (fn _=> [Asm_simp_tac 1]);
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
qed_goalw "insertE" Set.thy [insert_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   388
    "[| a : insert b A;  a=b ==> P;  a:A ==> P |] ==> P"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   389
 (fn major::prems=>
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   390
  [ (rtac (major RS UnE) 1),
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   391
    (REPEAT (eresolve_tac (prems @ [CollectE]) 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   392
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   393
(*Classical introduction rule*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   394
qed_goal "insertCI" Set.thy "(a~:B ==> a=b) ==> a: insert b B"
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   395
 (fn prems=>
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   396
  [ (Simp_tac 1),
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   397
    (REPEAT (ares_tac (prems@[disjCI]) 1)) ]);
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   398
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   399
AddSIs [insertCI]; 
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   400
AddSEs [insertE];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   401
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   402
section "Singletons, using insert";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   403
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   404
qed_goal "singletonI" Set.thy "a : {a}"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   405
 (fn _=> [ (rtac insertI1 1) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   406
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   407
goal Set.thy "!!a. b : {a} ==> b=a";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   408
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   409
qed "singletonD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   410
1776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   411
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
   412
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   413
qed_goal "singleton_iff" thy "(b : {a}) = (b=a)" 
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   414
(fn _ => [Blast_tac 1]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   415
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   416
goal Set.thy "!!a b. {a}={b} ==> a=b";
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2912
diff changeset
   417
by (blast_tac (!claset addEs [equalityE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   418
qed "singleton_inject";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   419
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   420
(*Redundant? But unlike insertCI, it proves the subgoal immediately!*)
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   421
AddSIs [singletonI];   
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   422
AddSDs [singleton_inject];
3718
d78cf498a88c Minor tidying to use Clarify_tac, etc.
paulson
parents: 3582
diff changeset
   423
AddSEs [singletonE];
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   424
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3718
diff changeset
   425
goal Set.thy "{x. x=a} = {a}";
3582
b87c86b6c291 Added {x.x=a} = a to !simpset.
nipkow
parents: 3469
diff changeset
   426
by(Blast_tac 1);
b87c86b6c291 Added {x.x=a} = a to !simpset.
nipkow
parents: 3469
diff changeset
   427
qed "singleton_conv";
b87c86b6c291 Added {x.x=a} = a to !simpset.
nipkow
parents: 3469
diff changeset
   428
Addsimps [singleton_conv];
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   429
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   430
section "The universal set -- UNIV";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   431
1882
67f49e8c4355 Proved bex_False
paulson
parents: 1841
diff changeset
   432
qed_goal "UNIV_I" Set.thy "x : UNIV"
67f49e8c4355 Proved bex_False
paulson
parents: 1841
diff changeset
   433
  (fn _ => [rtac ComplI 1, etac emptyE 1]);
67f49e8c4355 Proved bex_False
paulson
parents: 1841
diff changeset
   434
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   435
qed_goal "subset_UNIV" Set.thy "A <= UNIV"
1882
67f49e8c4355 Proved bex_False
paulson
parents: 1841
diff changeset
   436
  (fn _ => [rtac subsetI 1, rtac UNIV_I 1]);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   437
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   438
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   439
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
   440
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   441
goalw Set.thy [UNION_def] "(b: (UN x:A. B(x))) = (EX x:A. b: B(x))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   442
by (Blast_tac 1);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   443
qed "UN_iff";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   444
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   445
Addsimps [UN_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   446
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   447
(*The order of the premises presupposes that A is rigid; b may be flexible*)
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   448
goal Set.thy "!!b. [| a:A;  b: B(a) |] ==> b: (UN x:A. B(x))";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   449
by (Auto_tac());
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   450
qed "UN_I";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   451
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   452
val major::prems = goalw Set.thy [UNION_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   453
    "[| 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
   454
by (rtac (major RS CollectD RS bexE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   455
by (REPEAT (ares_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   456
qed "UN_E";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   457
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   458
AddIs  [UN_I];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   459
AddSEs [UN_E];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   460
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   461
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   462
    "[| A=B;  !!x. x:B ==> C(x) = D(x) |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   463
\    (UN x:A. C(x)) = (UN x:B. D(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   464
by (REPEAT (etac UN_E 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   465
     ORELSE ares_tac ([UN_I,equalityI,subsetI] @ 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
   466
                      (prems RL [equalityD1,equalityD2] RL [subsetD])) 1));
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   467
qed "UN_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   468
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   469
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   470
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
   471
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   472
goalw Set.thy [INTER_def] "(b: (INT x:A. B(x))) = (ALL x:A. b: B(x))";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   473
by (Auto_tac());
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   474
qed "INT_iff";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   475
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   476
Addsimps [INT_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   477
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   478
val prems = goalw Set.thy [INTER_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   479
    "(!!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
   480
by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   481
qed "INT_I";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   482
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   483
goal Set.thy "!!b. [| b : (INT x:A. B(x));  a:A |] ==> b: B(a)";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   484
by (Auto_tac());
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   485
qed "INT_D";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   486
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   487
(*"Classical" elimination -- by the Excluded Middle on a:A *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   488
val major::prems = goalw Set.thy [INTER_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   489
    "[| 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
   490
by (rtac (major RS CollectD RS ballE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   491
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   492
qed "INT_E";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   493
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   494
AddSIs [INT_I];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   495
AddEs  [INT_D, INT_E];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   496
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   497
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   498
    "[| A=B;  !!x. x:B ==> C(x) = D(x) |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   499
\    (INT x:A. C(x)) = (INT x:B. D(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   500
by (REPEAT_FIRST (resolve_tac [INT_I,equalityI,subsetI]));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   501
by (REPEAT (dtac INT_D 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   502
     ORELSE ares_tac (prems RL [equalityD1,equalityD2] RL [subsetD]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   503
qed "INT_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   504
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   505
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   506
section "Unions over a type; UNION1(B) = Union(range(B))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   507
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   508
goalw Set.thy [UNION1_def] "(b: (UN x. B(x))) = (EX x. b: B(x))";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   509
by (Simp_tac 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   510
by (Blast_tac 1);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   511
qed "UN1_iff";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   512
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   513
Addsimps [UN1_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   514
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   515
(*The order of the premises presupposes that A is rigid; b may be flexible*)
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   516
goal Set.thy "!!b. b: B(x) ==> b: (UN x. B(x))";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   517
by (Auto_tac());
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   518
qed "UN1_I";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   519
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   520
val major::prems = goalw Set.thy [UNION1_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   521
    "[| b : (UN x. B(x));  !!x. b: B(x) ==> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   522
by (rtac (major RS UN_E) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   523
by (REPEAT (ares_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   524
qed "UN1_E";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   525
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   526
AddIs  [UN1_I];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   527
AddSEs [UN1_E];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   528
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   529
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   530
section "Intersections over a type; INTER1(B) = Inter(range(B))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   531
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   532
goalw Set.thy [INTER1_def] "(b: (INT x. B(x))) = (ALL x. b: B(x))";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   533
by (Simp_tac 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   534
by (Blast_tac 1);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   535
qed "INT1_iff";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   536
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   537
Addsimps [INT1_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   538
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   539
val prems = goalw Set.thy [INTER1_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   540
    "(!!x. b: B(x)) ==> b : (INT x. B(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   541
by (REPEAT (ares_tac (INT_I::prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   542
qed "INT1_I";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   543
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   544
goal Set.thy "!!b. b : (INT x. B(x)) ==> b: B(a)";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   545
by (Asm_full_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   546
qed "INT1_D";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   547
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   548
AddSIs [INT1_I]; 
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   549
AddDs  [INT1_D];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   550
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   551
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   552
section "Union";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   553
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   554
goalw Set.thy [Union_def] "(A : Union(C)) = (EX X:C. A:X)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   555
by (Blast_tac 1);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   556
qed "Union_iff";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   557
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   558
Addsimps [Union_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   559
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   560
(*The order of the premises presupposes that C is rigid; A may be flexible*)
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   561
goal Set.thy "!!X. [| X:C;  A:X |] ==> A : Union(C)";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   562
by (Auto_tac());
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   563
qed "UnionI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   564
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   565
val major::prems = goalw Set.thy [Union_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   566
    "[| A : Union(C);  !!X.[| A:X;  X:C |] ==> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   567
by (rtac (major RS UN_E) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   568
by (REPEAT (ares_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   569
qed "UnionE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   570
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   571
AddIs  [UnionI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   572
AddSEs [UnionE];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   573
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   574
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   575
section "Inter";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   576
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   577
goalw Set.thy [Inter_def] "(A : Inter(C)) = (ALL X:C. A:X)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   578
by (Blast_tac 1);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   579
qed "Inter_iff";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   580
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   581
Addsimps [Inter_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   582
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   583
val prems = goalw Set.thy [Inter_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   584
    "[| !!X. X:C ==> A:X |] ==> A : Inter(C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   585
by (REPEAT (ares_tac ([INT_I] @ prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   586
qed "InterI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   587
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   588
(*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
   589
  A:X can hold when X:C does not!  This rule is analogous to "spec". *)
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   590
goal Set.thy "!!X. [| A : Inter(C);  X:C |] ==> A:X";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   591
by (Auto_tac());
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   592
qed "InterD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   593
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   594
(*"Classical" elimination rule -- does not require proving X:C *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   595
val major::prems = goalw Set.thy [Inter_def]
2721
f08042e18c7d New version of InterE, like its ZF counterpart
paulson
parents: 2608
diff changeset
   596
    "[| A : Inter(C);  X~:C ==> R;  A:X ==> R |] ==> R";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   597
by (rtac (major RS INT_E) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   598
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   599
qed "InterE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   600
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   601
AddSIs [InterI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   602
AddEs  [InterD, InterE];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   603
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   604
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   605
(*** Image of a set under a function ***)
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   606
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   607
(*Frequently b does not have the syntactic form of f(x).*)
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   608
val prems = goalw thy [image_def] "[| b=f(x);  x:A |] ==> b : f``A";
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   609
by (REPEAT (resolve_tac (prems @ [CollectI,bexI,prem]) 1));
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   610
qed "image_eqI";
3909
e48e2fb8b895 Added image_eqI to simpset.
nipkow
parents: 3905
diff changeset
   611
Addsimps [image_eqI];
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   612
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   613
bind_thm ("imageI", refl RS image_eqI);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   614
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   615
(*The eta-expansion gives variable-name preservation.*)
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   616
val major::prems = goalw thy [image_def]
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3718
diff changeset
   617
    "[| b : (%x. f(x))``A;  !!x.[| b=f(x);  x:A |] ==> P |] ==> P"; 
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   618
by (rtac (major RS CollectD RS bexE) 1);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   619
by (REPEAT (ares_tac prems 1));
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   620
qed "imageE";
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   621
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   622
AddIs  [image_eqI];
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   623
AddSEs [imageE]; 
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   624
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   625
goalw thy [o_def] "(f o g)``r = f``(g``r)";
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2912
diff changeset
   626
by (Blast_tac 1);
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   627
qed "image_compose";
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   628
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   629
goal thy "f``(A Un B) = f``A Un f``B";
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2912
diff changeset
   630
by (Blast_tac 1);
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   631
qed "image_Un";
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   632
3960
7a38fae985f9 New rewrite rules image_iff
paulson
parents: 3919
diff changeset
   633
goal Set.thy "(z : f``A) = (EX x:A. z = f x)";
7a38fae985f9 New rewrite rules image_iff
paulson
parents: 3919
diff changeset
   634
by (Blast_tac 1);
7a38fae985f9 New rewrite rules image_iff
paulson
parents: 3919
diff changeset
   635
qed "image_iff";
7a38fae985f9 New rewrite rules image_iff
paulson
parents: 3919
diff changeset
   636
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   637
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   638
(*** Range of a function -- just a translation for image! ***)
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   639
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   640
goal thy "!!b. b=f(x) ==> b : range(f)";
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   641
by (EVERY1 [etac image_eqI, rtac UNIV_I]);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   642
bind_thm ("range_eqI", UNIV_I RSN (2,image_eqI));
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   643
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   644
bind_thm ("rangeI", UNIV_I RS imageI);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   645
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   646
val [major,minor] = goal thy 
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3718
diff changeset
   647
    "[| b : range(%x. f(x));  !!x. b=f(x) ==> P |] ==> P"; 
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   648
by (rtac (major RS imageE) 1);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   649
by (etac minor 1);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   650
qed "rangeE";
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   651
1776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   652
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   653
(*** Set reasoning tools ***)
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   654
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   655
3912
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   656
(** Rewrite rules for boolean case-splitting: faster than 
3919
c036caebfc75 setloop split_tac -> addsplits
nipkow
parents: 3912
diff changeset
   657
	addsplits[expand_if]
3912
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   658
**)
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   659
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   660
bind_thm ("expand_if_eq1", read_instantiate [("P", "%x. x = ?b")] expand_if);
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   661
bind_thm ("expand_if_eq2", read_instantiate [("P", "%x. ?a = x")] expand_if);
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   662
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   663
bind_thm ("expand_if_mem1", 
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   664
    read_instantiate_sg (sign_of Set.thy) [("P", "%x. x : ?b")] expand_if);
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   665
bind_thm ("expand_if_mem2", 
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   666
    read_instantiate_sg (sign_of Set.thy) [("P", "%x. ?a : x")] expand_if);
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   667
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   668
val expand_ifs = [if_bool_eq, expand_if_eq1, expand_if_eq2,
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   669
		  expand_if_mem1, expand_if_mem2];
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   670
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   671
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   672
(*Each of these has ALREADY been added to !simpset above.*)
2024
909153d8318f Rationalized the rewriting of membership for {} and insert
paulson
parents: 1985
diff changeset
   673
val mem_simps = [insert_iff, empty_iff, Un_iff, Int_iff, Compl_iff, Diff_iff, 
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   674
                 mem_Collect_eq, 
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   675
		 UN_iff, UN1_iff, Union_iff, 
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   676
		 INT_iff, INT1_iff, Inter_iff];
1776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   677
1937
e59ff0eb1e91 Proved mem_if
paulson
parents: 1920
diff changeset
   678
(*Not for Addsimps -- it can cause goals to blow up!*)
e59ff0eb1e91 Proved mem_if
paulson
parents: 1920
diff changeset
   679
goal Set.thy "(a : (if Q then x else y)) = ((Q --> a:x) & (~Q --> a : y))";
3919
c036caebfc75 setloop split_tac -> addsplits
nipkow
parents: 3912
diff changeset
   680
by (simp_tac (!simpset addsplits [expand_if]) 1);
1937
e59ff0eb1e91 Proved mem_if
paulson
parents: 1920
diff changeset
   681
qed "mem_if";
e59ff0eb1e91 Proved mem_if
paulson
parents: 1920
diff changeset
   682
1776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   683
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
   684
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   685
simpset := !simpset addcongs [ball_cong,bex_cong]
1776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   686
                    setmksimps (mksimps mksimps_pairs);
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   687
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   688
Addsimps[subset_UNIV, empty_subsetI, subset_refl];
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   689
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   690
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   691
(*** < ***)
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   692
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   693
goalw Set.thy [psubset_def] "!!A::'a set. [| A <= B; A ~= B |] ==> A<B";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   694
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   695
qed "psubsetI";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   696
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   697
goalw Set.thy [psubset_def]
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   698
    "!!x. A < insert x B ==> (x ~: A) & A<=B | x:A & A-{x}<B";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   699
by (Auto_tac());
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   700
qed "psubset_insertD";
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   701
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   702
bind_thm ("psubset_eq", psubset_def RS meta_eq_to_obj_eq);