src/HOL/equalities.ML
author wenzelm
Mon, 03 Nov 1997 12:13:18 +0100
changeset 4089 96fba19bcbe2
parent 4059 59c1422c9da5
child 4136 ba267836dd7a
permissions -rw-r--r--
isatool fixclasimp;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
     1
(*  Title:      HOL/equalities
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
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   1994  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
Equalities involving union, intersection, inclusion, etc.
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
writeln"File HOL/equalities";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
1754
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1748
diff changeset
    11
AddSIs [equalityI];
852093aeb0ab Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1748
diff changeset
    12
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    13
section "{}";
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    14
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    15
goal thy "{x. False} = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    16
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    17
qed "Collect_False_empty";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    18
Addsimps [Collect_False_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    19
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    20
goal thy "(A <= {}) = (A = {})";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    21
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    22
qed "subset_empty";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    23
Addsimps [subset_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    24
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    25
goalw thy [psubset_def] "~ (A < {})";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    26
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    27
qed "not_psubset_empty";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    28
AddIffs [not_psubset_empty];
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    29
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    30
section "insert";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    32
(*NOT SUITABLE FOR REWRITING since {a} == insert a {}*)
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    33
goal thy "insert a A = {a} Un A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    34
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    35
qed "insert_is_Un";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    36
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    37
goal thy "insert a A ~= {}";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
    38
by (blast_tac (claset() addEs [equalityCE]) 1);
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
    39
qed"insert_not_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    40
Addsimps[insert_not_empty];
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
    41
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
    42
bind_thm("empty_not_insert",insert_not_empty RS not_sym);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    43
Addsimps[empty_not_insert];
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
    44
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    45
goal thy "!!a. a:A ==> insert a A = A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    46
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
qed "insert_absorb";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    49
goal thy "insert x (insert x A) = insert x A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    50
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    51
qed "insert_absorb2";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    52
Addsimps [insert_absorb2];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    53
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    54
goal thy "insert x (insert y A) = insert y (insert x A)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    55
by (Blast_tac 1);
1879
83c70ad381c1 Added insert_commute
paulson
parents: 1843
diff changeset
    56
qed "insert_commute";
83c70ad381c1 Added insert_commute
paulson
parents: 1843
diff changeset
    57
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    58
goal thy "(insert x A <= B) = (x:B & A <= B)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    59
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
qed "insert_subset";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    61
Addsimps[insert_subset];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    62
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    63
goal thy "!!a. insert a A ~= insert a B ==> A ~= B";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    64
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    65
qed "insert_lim";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    66
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    67
(* use new B rather than (A-{a}) to avoid infinite unfolding *)
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    68
goal thy "!!a. a:A ==> ? B. A = insert a B & a ~: B";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    69
by (res_inst_tac [("x","A-{a}")] exI 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    70
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    71
qed "mk_disjoint_insert";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    72
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    73
goal thy
1843
a6d7aef48c2f Removed the unused eq_cs, and added some distributive laws
paulson
parents: 1786
diff changeset
    74
    "!!A. A~={} ==> (UN x:A. insert a (B x)) = insert a (UN x:A. B x)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    75
by (Blast_tac 1);
1843
a6d7aef48c2f Removed the unused eq_cs, and added some distributive laws
paulson
parents: 1786
diff changeset
    76
qed "UN_insert_distrib";
a6d7aef48c2f Removed the unused eq_cs, and added some distributive laws
paulson
parents: 1786
diff changeset
    77
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    78
goal thy "(UN x. insert a (B x)) = insert a (UN x. B x)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    79
by (Blast_tac 1);
1843
a6d7aef48c2f Removed the unused eq_cs, and added some distributive laws
paulson
parents: 1786
diff changeset
    80
qed "UN1_insert_distrib";
a6d7aef48c2f Removed the unused eq_cs, and added some distributive laws
paulson
parents: 1786
diff changeset
    81
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
    82
section "``";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    84
goal thy "f``{} = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    85
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
qed "image_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    87
Addsimps[image_empty];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    89
goal thy "f``insert a B = insert (f a) (f``B)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
    90
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
qed "image_insert";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    92
Addsimps[image_insert];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    94
goal thy  "(f `` (UNION A B)) = (UN x:A.(f `` (B x)))";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    95
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    96
qed "image_UNION";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    97
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
    98
goal thy "(%x. x) `` Y = Y";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
    99
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   100
qed "image_id";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   101
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   102
goal thy "f``(g``A) = (%x. f (g x)) `` A";
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3426
diff changeset
   103
by (Blast_tac 1);
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   104
qed "image_image";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   105
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   106
qed_goal "ball_image" Set.thy "(!y:F``S. P y) = (!x:S. P (F x))"
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   107
 (fn _ => [Blast_tac 1]);
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   108
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   109
goal thy "!!x. x:A ==> insert (f x) (f``A) = f``A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   110
by (Blast_tac 1);
1884
5a1f81da3e12 Proved insert_image
paulson
parents: 1879
diff changeset
   111
qed "insert_image";
5a1f81da3e12 Proved insert_image
paulson
parents: 1879
diff changeset
   112
Addsimps [insert_image];
5a1f81da3e12 Proved insert_image
paulson
parents: 1879
diff changeset
   113
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   114
goal thy "(f``A = {}) = (A = {})";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   115
by (blast_tac (claset() addSEs [equalityE]) 1);
3415
c068bd2f0bbd Moved image_is_empty from Finite.ML to equalities.ML
nipkow
parents: 3384
diff changeset
   116
qed "image_is_empty";
c068bd2f0bbd Moved image_is_empty from Finite.ML to equalities.ML
nipkow
parents: 3384
diff changeset
   117
AddIffs [image_is_empty];
c068bd2f0bbd Moved image_is_empty from Finite.ML to equalities.ML
nipkow
parents: 3384
diff changeset
   118
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   119
goalw thy [image_def]
1763
fb07e359b59f expanded TABs
berghofe
parents: 1754
diff changeset
   120
"(%x. if P x then f x else g x) `` S                    \
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   121
\ = (f `` ({x. x:S & P x})) Un (g `` ({x. x:S & ~(P x)}))";
2031
03a843f0f447 Ran expandshort
paulson
parents: 2024
diff changeset
   122
by (split_tac [expand_if] 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   123
by (Blast_tac 1);
1748
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
   124
qed "if_image_distrib";
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
   125
Addsimps[if_image_distrib];
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
   126
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
   127
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   128
section "Int";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   130
goal thy "A Int A = A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   131
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
qed "Int_absorb";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   133
Addsimps[Int_absorb];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   135
goal thy "A Int B  =  B Int A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   136
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
qed "Int_commute";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   138
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   139
goal thy "(A Int B) Int C  =  A Int (B Int C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   140
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   141
qed "Int_assoc";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   142
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   143
goal thy "{} Int B = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   144
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   145
qed "Int_empty_left";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   146
Addsimps[Int_empty_left];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   147
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   148
goal thy "A Int {} = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   149
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   150
qed "Int_empty_right";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   151
Addsimps[Int_empty_right];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   152
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   153
goal thy "(A Int B = {}) = (A <= Compl B)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   154
by (blast_tac (claset() addSEs [equalityE]) 1);
3356
9b899eb8a036 New theorem disjoint_eq_subset_Compl
paulson
parents: 3348
diff changeset
   155
qed "disjoint_eq_subset_Compl";
9b899eb8a036 New theorem disjoint_eq_subset_Compl
paulson
parents: 3348
diff changeset
   156
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   157
goal thy "UNIV Int B = B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   158
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   159
qed "Int_UNIV_left";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   160
Addsimps[Int_UNIV_left];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   161
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   162
goal thy "A Int UNIV = A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   163
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   164
qed "Int_UNIV_right";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   165
Addsimps[Int_UNIV_right];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   166
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   167
goal thy "A Int (B Un C)  =  (A Int B) Un (A Int C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   168
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   169
qed "Int_Un_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   170
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   171
goal thy "(B Un C) Int A =  (B Int A) Un (C Int A)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   172
by (Blast_tac 1);
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1564
diff changeset
   173
qed "Int_Un_distrib2";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1564
diff changeset
   174
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   175
goal thy "(A<=B) = (A Int B = A)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   176
by (blast_tac (claset() addSEs [equalityE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
qed "subset_Int_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   179
goal thy "(A Int B = UNIV) = (A = UNIV & B = UNIV)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   180
by (blast_tac (claset() addEs [equalityCE]) 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   181
qed "Int_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   182
Addsimps[Int_UNIV];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   183
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   184
section "Un";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   185
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   186
goal thy "A Un A = A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   187
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   188
qed "Un_absorb";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   189
Addsimps[Un_absorb];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   190
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   191
goal thy " A Un (A Un B) = A Un B";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   192
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   193
qed "Un_left_absorb";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   194
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   195
goal thy "A Un B  =  B Un A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   196
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   197
qed "Un_commute";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   199
goal thy " A Un (B Un C) = B Un (A Un C)";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   200
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   201
qed "Un_left_commute";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   202
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   203
goal thy "(A Un B) Un C  =  A Un (B Un C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   204
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   205
qed "Un_assoc";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   206
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   207
goal thy "{} Un B = B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   208
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   209
qed "Un_empty_left";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   210
Addsimps[Un_empty_left];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   211
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   212
goal thy "A Un {} = A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   213
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   214
qed "Un_empty_right";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   215
Addsimps[Un_empty_right];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   216
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   217
goal thy "UNIV Un B = UNIV";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   218
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   219
qed "Un_UNIV_left";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   220
Addsimps[Un_UNIV_left];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   221
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   222
goal thy "A Un UNIV = UNIV";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   223
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   224
qed "Un_UNIV_right";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   225
Addsimps[Un_UNIV_right];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   226
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   227
goal thy "(insert a B) Un C = insert a (B Un C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   228
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   229
qed "Un_insert_left";
3384
5ef99c94e1fb Now Un_insert_left, Un_insert_right are default rewrite rules
paulson
parents: 3356
diff changeset
   230
Addsimps[Un_insert_left];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   231
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   232
goal thy "A Un (insert a B) = insert a (A Un B)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   233
by (Blast_tac 1);
1917
27b71d839d50 Added proof of Un_insert_right
paulson
parents: 1884
diff changeset
   234
qed "Un_insert_right";
3384
5ef99c94e1fb Now Un_insert_left, Un_insert_right are default rewrite rules
paulson
parents: 3356
diff changeset
   235
Addsimps[Un_insert_right];
1917
27b71d839d50 Added proof of Un_insert_right
paulson
parents: 1884
diff changeset
   236
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   237
goal thy "(insert a B) Int C = (if a:C then insert a (B Int C) \
3356
9b899eb8a036 New theorem disjoint_eq_subset_Compl
paulson
parents: 3348
diff changeset
   238
\                                          else        B Int C)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   239
by (simp_tac (simpset() addsplits [expand_if]) 1);
3356
9b899eb8a036 New theorem disjoint_eq_subset_Compl
paulson
parents: 3348
diff changeset
   240
by (Blast_tac 1);
9b899eb8a036 New theorem disjoint_eq_subset_Compl
paulson
parents: 3348
diff changeset
   241
qed "Int_insert_left";
9b899eb8a036 New theorem disjoint_eq_subset_Compl
paulson
parents: 3348
diff changeset
   242
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   243
goal thy "A Int (insert a B) = (if a:A then insert a (A Int B) \
3356
9b899eb8a036 New theorem disjoint_eq_subset_Compl
paulson
parents: 3348
diff changeset
   244
\                                          else        A Int B)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   245
by (simp_tac (simpset() addsplits [expand_if]) 1);
3356
9b899eb8a036 New theorem disjoint_eq_subset_Compl
paulson
parents: 3348
diff changeset
   246
by (Blast_tac 1);
9b899eb8a036 New theorem disjoint_eq_subset_Compl
paulson
parents: 3348
diff changeset
   247
qed "Int_insert_right";
9b899eb8a036 New theorem disjoint_eq_subset_Compl
paulson
parents: 3348
diff changeset
   248
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   249
goal thy "(A Int B) Un C  =  (A Un C) Int (B Un C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   250
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   251
qed "Un_Int_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   252
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   253
goal thy
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   254
 "(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   255
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   256
qed "Un_Int_crazy";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   257
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   258
goal thy "(A<=B) = (A Un B = B)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   259
by (blast_tac (claset() addSEs [equalityE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   260
qed "subset_Un_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   261
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   262
goal thy "(A <= insert b C) = (A <= C | b:A & A-{b} <= C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   263
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   264
qed "subset_insert_iff";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   265
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   266
goal thy "(A Un B = {}) = (A = {} & B = {})";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   267
by (blast_tac (claset() addEs [equalityCE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   268
qed "Un_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   269
Addsimps[Un_empty];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   270
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   271
section "Compl";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   272
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   273
goal thy "A Int Compl(A) = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   274
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   275
qed "Compl_disjoint";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   276
Addsimps[Compl_disjoint];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   277
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   278
goal thy "A Un Compl(A) = UNIV";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   279
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   280
qed "Compl_partition";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   281
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   282
goal thy "Compl(Compl(A)) = A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   283
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   284
qed "double_complement";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   285
Addsimps[double_complement];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   286
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   287
goal thy "Compl(A Un B) = Compl(A) Int Compl(B)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   288
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   289
qed "Compl_Un";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   290
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   291
goal thy "Compl(A Int B) = Compl(A) Un Compl(B)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   292
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   293
qed "Compl_Int";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   294
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   295
goal thy "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   296
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   297
qed "Compl_UN";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   298
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   299
goal thy "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   300
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   301
qed "Compl_INT";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   302
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   303
(*Halmos, Naive Set Theory, page 16.*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   304
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   305
goal thy "((A Int B) Un C = A Int (B Un C)) = (C<=A)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   306
by (blast_tac (claset() addSEs [equalityE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   307
qed "Un_Int_assoc_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   308
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   309
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   310
section "Union";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   311
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   312
goal thy "Union({}) = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   313
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   314
qed "Union_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   315
Addsimps[Union_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   316
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   317
goal thy "Union(UNIV) = UNIV";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   318
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   319
qed "Union_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   320
Addsimps[Union_UNIV];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   321
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   322
goal thy "Union(insert a B) = a Un Union(B)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   323
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   324
qed "Union_insert";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   325
Addsimps[Union_insert];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   326
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   327
goal thy "Union(A Un B) = Union(A) Un Union(B)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   328
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   329
qed "Union_Un_distrib";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   330
Addsimps[Union_Un_distrib];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   331
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   332
goal thy "Union(A Int B) <= Union(A) Int Union(B)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   333
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   334
qed "Union_Int_subset";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   335
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   336
goal thy "(Union M = {}) = (! A : M. A = {})"; 
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   337
by (blast_tac (claset() addEs [equalityE]) 1);
4003
nipkow
parents: 3919
diff changeset
   338
qed"Union_empty_conv"; 
nipkow
parents: 3919
diff changeset
   339
AddIffs [Union_empty_conv];
nipkow
parents: 3919
diff changeset
   340
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   341
val prems = goal thy
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   342
   "(Union(C) Int A = {}) = (! B:C. B Int A = {})";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   343
by (blast_tac (claset() addSEs [equalityE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   344
qed "Union_disjoint";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   345
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   346
section "Inter";
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   347
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   348
goal thy "Inter({}) = UNIV";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   349
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   350
qed "Inter_empty";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   351
Addsimps[Inter_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   352
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   353
goal thy "Inter(UNIV) = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   354
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   355
qed "Inter_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   356
Addsimps[Inter_UNIV];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   357
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   358
goal thy "Inter(insert a B) = a Int Inter(B)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   359
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   360
qed "Inter_insert";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   361
Addsimps[Inter_insert];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   362
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   363
goal thy "Inter(A) Un Inter(B) <= Inter(A Int B)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   364
by (Blast_tac 1);
1564
822575c737bd Deleted faulty comment; proved new rule Inter_Un_subset
paulson
parents: 1553
diff changeset
   365
qed "Inter_Un_subset";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   366
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   367
goal thy "Inter(A Un B) = Inter(A) Int Inter(B)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   368
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   369
qed "Inter_Un_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   370
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   371
section "UN and INT";
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
(*Basic identities*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   374
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   375
goal thy "(UN x:{}. B x) = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   376
by (Blast_tac 1);
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   377
qed "UN_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   378
Addsimps[UN_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   379
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   380
goal thy "(UN x:A. {}) = {}";
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3426
diff changeset
   381
by (Blast_tac 1);
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   382
qed "UN_empty2";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   383
Addsimps[UN_empty2];
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   384
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   385
goal thy "(UN x:UNIV. B x) = (UN x. B x)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   386
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   387
qed "UN_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   388
Addsimps[UN_UNIV];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   389
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   390
goal thy "(INT x:{}. B x) = UNIV";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   391
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   392
qed "INT_empty";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   393
Addsimps[INT_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   394
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   395
goal thy "(INT x:UNIV. B x) = (INT x. B x)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   396
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   397
qed "INT_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   398
Addsimps[INT_UNIV];
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   399
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   400
goal thy "(UN x:insert a A. B x) = B a Un UNION A B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   401
by (Blast_tac 1);
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   402
qed "UN_insert";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   403
Addsimps[UN_insert];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   404
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   405
goal thy "(UN i: A Un B. M i) = ((UN i: A. M i) Un (UN i:B. M i))";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   406
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   407
qed "UN_Un";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   408
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   409
goal thy "(INT x:insert a A. B x) = B a Int INTER A B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   410
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   411
qed "INT_insert";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   412
Addsimps[INT_insert];
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   413
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   414
goal thy
2021
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   415
    "!!A. A~={} ==> (INT x:A. insert a (B x)) = insert a (INT x:A. B x)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   416
by (Blast_tac 1);
2021
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   417
qed "INT_insert_distrib";
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   418
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   419
goal thy "(INT x. insert a (B x)) = insert a (INT x. B x)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   420
by (Blast_tac 1);
2021
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   421
qed "INT1_insert_distrib";
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   422
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   423
goal thy "Union(B``A) = (UN x:A. B(x))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   424
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   425
qed "Union_image_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   426
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   427
goal thy "Inter(B``A) = (INT x:A. B(x))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   428
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   429
qed "Inter_image_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   430
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   431
goal thy "!!A. a: A ==> (UN y:A. c) = c";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   432
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   433
qed "UN_constant";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   434
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   435
goal thy "!!A. a: A ==> (INT y:A. c) = c";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   436
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   437
qed "INT_constant";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   438
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   439
goal thy "(UN x. B) = B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   440
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   441
qed "UN1_constant";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   442
Addsimps[UN1_constant];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   443
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   444
goal thy "(INT x. B) = B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   445
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   446
qed "INT1_constant";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   447
Addsimps[INT1_constant];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   448
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   449
goal thy "(UN x:A. B(x)) = Union({Y. ? x:A. Y=B(x)})";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   450
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   451
qed "UN_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   452
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   453
(*Look: it has an EXISTENTIAL quantifier*)
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   454
goal thy "(INT x:A. B(x)) = Inter({Y. ? x:A. Y=B(x)})";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   455
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   456
qed "INT_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   457
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   458
goalw thy [o_def] "UNION A (g o f) = UNION (f``A) g";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   459
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   460
qed "UNION_o";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   461
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   462
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   463
(*Distributive laws...*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   464
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   465
goal thy "A Int Union(B) = (UN C:B. A Int C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   466
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   467
qed "Int_Union";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   468
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   469
(* Devlin, Setdamentals of Contemporary Set Theory, page 12, exercise 5: 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   470
   Union of a family of unions **)
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   471
goal thy "(UN x:C. A(x) Un B(x)) = Union(A``C)  Un  Union(B``C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   472
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   473
qed "Un_Union_image";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   474
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   475
(*Equivalent version*)
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   476
goal thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i))  Un  (UN i:I. B(i))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   477
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   478
qed "UN_Un_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   479
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   480
goal thy "A Un Inter(B) = (INT C:B. A Un C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   481
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   482
qed "Un_Inter";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   483
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   484
goal thy "(INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   485
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   486
qed "Int_Inter_image";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   487
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   488
(*Equivalent version*)
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   489
goal thy "(INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   490
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   491
qed "INT_Int_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   492
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   493
(*Halmos, Naive Set Theory, page 35.*)
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   494
goal thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   495
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   496
qed "Int_UN_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   497
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   498
goal thy "B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   499
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   500
qed "Un_INT_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   501
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   502
goal thy
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   503
    "(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   504
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   505
qed "Int_UN_distrib2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   506
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   507
goal thy
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   508
    "(INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   509
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   510
qed "Un_INT_distrib2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   511
2512
0231e4f467f2 Got rid of Alls in List.
nipkow
parents: 2031
diff changeset
   512
0231e4f467f2 Got rid of Alls in List.
nipkow
parents: 2031
diff changeset
   513
section"Bounded quantifiers";
0231e4f467f2 Got rid of Alls in List.
nipkow
parents: 2031
diff changeset
   514
3860
a29ab43f7174 More lemmas, esp. ~Bex and ~Ball conversions.
nipkow
parents: 3842
diff changeset
   515
(** The following are not added to the default simpset because
a29ab43f7174 More lemmas, esp. ~Bex and ~Ball conversions.
nipkow
parents: 3842
diff changeset
   516
    (a) they duplicate the body and (b) there are no similar rules for Int. **)
2512
0231e4f467f2 Got rid of Alls in List.
nipkow
parents: 2031
diff changeset
   517
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   518
goal thy "(ALL x:A Un B. P x) = ((ALL x:A. P x) & (ALL x:B. P x))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   519
by (Blast_tac 1);
2519
761e3094e32f New rewrites for bounded quantifiers
paulson
parents: 2513
diff changeset
   520
qed "ball_Un";
761e3094e32f New rewrites for bounded quantifiers
paulson
parents: 2513
diff changeset
   521
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   522
goal thy "(EX x:A Un B. P x) = ((EX x:A. P x) | (EX x:B. P x))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   523
by (Blast_tac 1);
2519
761e3094e32f New rewrites for bounded quantifiers
paulson
parents: 2513
diff changeset
   524
qed "bex_Un";
2512
0231e4f467f2 Got rid of Alls in List.
nipkow
parents: 2031
diff changeset
   525
0231e4f467f2 Got rid of Alls in List.
nipkow
parents: 2031
diff changeset
   526
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   527
section "-";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   528
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   529
goal thy "A-A = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   530
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   531
qed "Diff_cancel";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   532
Addsimps[Diff_cancel];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   533
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   534
goal thy "{}-A = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   535
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   536
qed "empty_Diff";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   537
Addsimps[empty_Diff];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   538
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   539
goal thy "A-{} = A";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   540
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   541
qed "Diff_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   542
Addsimps[Diff_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   543
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   544
goal thy "A-UNIV = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   545
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   546
qed "Diff_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   547
Addsimps[Diff_UNIV];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   548
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   549
goal thy "!!x. x~:A ==> A - insert x B = A-B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   550
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   551
qed "Diff_insert0";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   552
Addsimps [Diff_insert0];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   553
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   554
(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*)
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   555
goal thy "A - insert a B = A - B - {a}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   556
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   557
qed "Diff_insert";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   558
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   559
(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*)
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   560
goal thy "A - insert a B = A - {a} - B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   561
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   562
qed "Diff_insert2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   563
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   564
goal thy "insert x A - B = (if x:B then A-B else insert x (A-B))";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   565
by (simp_tac (simpset() addsplits [expand_if]) 1);
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   566
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   567
qed "insert_Diff_if";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   568
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   569
goal thy "!!x. x:B ==> insert x A - B = A-B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   570
by (Blast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   571
qed "insert_Diff1";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   572
Addsimps [insert_Diff1];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   573
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   574
goal thy "!!a. a:A ==> insert a (A-{a}) = A";
2922
580647a879cf Using Blast_tac
paulson
parents: 2912
diff changeset
   575
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   576
qed "insert_Diff";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   577
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   578
goal thy "A Int (B-A) = {}";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   579
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   580
qed "Diff_disjoint";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   581
Addsimps[Diff_disjoint];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   582
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   583
goal thy "!!A. A<=B ==> A Un (B-A) = B";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   584
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   585
qed "Diff_partition";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   586
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   587
goal thy "!!A. [| A<=B; B<= C |] ==> (B - (C - A)) = (A :: 'a set)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   588
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   589
qed "double_diff";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   590
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   591
goal thy "A - (B Un C) = (A-B) Int (A-C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   592
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   593
qed "Diff_Un";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   594
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   595
goal thy "A - (B Int C) = (A-B) Un (A-C)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2519
diff changeset
   596
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   597
qed "Diff_Int";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   598
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   599
goal thy "(A Un B) - C = (A - C) Un (B - C)";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   600
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   601
qed "Un_Diff";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   602
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   603
goal thy "(A Int B) - C = (A - C) Int (B - C)";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   604
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   605
qed "Int_Diff";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   606
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   607
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   608
section "Miscellany";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   609
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   610
goal thy "(A = B) = ((A <= (B::'a set)) & (B<=A))";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   611
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   612
qed "set_eq_subset";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   613
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   614
goal thy "A <= B =  (! t. t:A --> t:B)";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   615
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   616
qed "subset_iff";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   617
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   618
goalw thy [psubset_def] "((A::'a set) <= B) = ((A < B) | (A=B))";
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   619
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2922
diff changeset
   620
qed "subset_iff_psubset_eq";
2021
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   621
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   622
goal thy "(!x. x ~: A) = (A={})";
3896
ee8ebb74ec00 Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents: 3860
diff changeset
   623
by(Blast_tac 1);
ee8ebb74ec00 Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents: 3860
diff changeset
   624
qed "all_not_in_conv";
3907
51c00e87bd6e AddIffs [all_not_in_conv];
nipkow
parents: 3896
diff changeset
   625
AddIffs [all_not_in_conv];
3896
ee8ebb74ec00 Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents: 3860
diff changeset
   626
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   627
goalw thy [Pow_def] "Pow {} = {{}}";
3348
3f9a806f061e Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   628
by (Auto_tac());
3f9a806f061e Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   629
qed "Pow_empty";
3f9a806f061e Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   630
Addsimps [Pow_empty];
3f9a806f061e Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   631
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   632
goal thy "Pow (insert a A) = Pow A Un (insert a `` Pow A)";
3724
f33e301a89f5 Step_tac -> Safe_tac
paulson
parents: 3457
diff changeset
   633
by Safe_tac;
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3426
diff changeset
   634
by (etac swap 1);
3348
3f9a806f061e Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   635
by (res_inst_tac [("x", "x-{a}")] image_eqI 1);
3f9a806f061e Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   636
by (ALLGOALS Blast_tac);
3f9a806f061e Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   637
qed "Pow_insert";
3f9a806f061e Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents: 3222
diff changeset
   638
2021
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   639
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   640
(** Miniscoping: pushing in big Unions and Intersections **)
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   641
local
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 4003
diff changeset
   642
  fun prover s = prove_goal thy s (fn _ => [Blast_tac 1])
2021
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   643
in
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   644
val UN1_simps = map prover 
2031
03a843f0f447 Ran expandshort
paulson
parents: 2024
diff changeset
   645
                ["(UN x. insert a (B x)) = insert a (UN x. B x)",
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   646
                 "(UN x. A x Int B)  = ((UN x. A x) Int B)",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   647
                 "(UN x. A Int B x)  = (A Int (UN x. B x))",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   648
                 "(UN x. A x Un B)   = ((UN x. A x) Un B)",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   649
                 "(UN x. A Un B x)   = (A Un (UN x. B x))",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   650
                 "(UN x. A x - B)    = ((UN x. A x) - B)",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   651
                 "(UN x. A - B x)    = (A - (INT x. B x))"];
2021
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   652
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   653
val INT1_simps = map prover
2031
03a843f0f447 Ran expandshort
paulson
parents: 2024
diff changeset
   654
                ["(INT x. insert a (B x)) = insert a (INT x. B x)",
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   655
                 "(INT x. A x Int B) = ((INT x. A x) Int B)",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   656
                 "(INT x. A Int B x) = (A Int (INT x. B x))",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   657
                 "(INT x. A x Un B)  = ((INT x. A x) Un B)",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   658
                 "(INT x. A Un B x)  = (A Un (INT x. B x))",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   659
                 "(INT x. A x - B)   = ((INT x. A x) - B)",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   660
                 "(INT x. A - B x)   = (A - (UN x. B x))"];
2021
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   661
2513
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   662
val UN_simps = map prover 
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   663
                ["(UN x:C. A x Int B)  = ((UN x:C. A x) Int B)",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   664
                 "(UN x:C. A Int B x)  = (A Int (UN x:C. B x))",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   665
                 "(UN x:C. A x - B)    = ((UN x:C. A x) - B)",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   666
                 "(UN x:C. A - B x)    = (A - (INT x:C. B x))"];
2513
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   667
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   668
val INT_simps = map prover
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   669
                ["(INT x:C. insert a (B x)) = insert a (INT x:C. B x)",
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   670
                 "(INT x:C. A x Un B)  = ((INT x:C. A x) Un B)",
b55686a7b22c fixed dots;
wenzelm
parents: 3724
diff changeset
   671
                 "(INT x:C. A Un B x)  = (A Un (INT x:C. B x))"];
2513
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   672
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   673
(*The missing laws for bounded Unions and Intersections are conditional
2021
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   674
  on the index set's being non-empty.  Thus they are probably NOT worth 
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   675
  adding as default rewrites.*)
2513
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   676
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   677
val ball_simps = map prover
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   678
    ["(ALL x:A. P x | Q) = ((ALL x:A. P x) | Q)",
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   679
     "(ALL x:A. P | Q x) = (P | (ALL x:A. Q x))",
3422
16ae2c20801c New miniscoping rules for ALL
paulson
parents: 3415
diff changeset
   680
     "(ALL x:A. P --> Q x) = (P --> (ALL x:A. Q x))",
16ae2c20801c New miniscoping rules for ALL
paulson
parents: 3415
diff changeset
   681
     "(ALL x:A. P x --> Q) = ((EX x:A. P x) --> Q)",
2513
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   682
     "(ALL x:{}. P x) = True",
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   683
     "(ALL x:insert a B. P x) = (P(a) & (ALL x:B. P x))",
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   684
     "(ALL x:Union(A). P x) = (ALL y:A. ALL x:y. P x)",
3860
a29ab43f7174 More lemmas, esp. ~Bex and ~Ball conversions.
nipkow
parents: 3842
diff changeset
   685
     "(ALL x:Collect Q. P x) = (ALL x. Q x --> P x)",
a29ab43f7174 More lemmas, esp. ~Bex and ~Ball conversions.
nipkow
parents: 3842
diff changeset
   686
     "(ALL x:f``A. P x) = (ALL x:A. P(f x))",
a29ab43f7174 More lemmas, esp. ~Bex and ~Ball conversions.
nipkow
parents: 3842
diff changeset
   687
     "(~(ALL x:A. P x)) = (EX x:A. ~P x)"];
2513
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   688
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   689
val ball_conj_distrib = 
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   690
    prover "(ALL x:A. P x & Q x) = ((ALL x:A. P x) & (ALL x:A. Q x))";
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   691
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   692
val bex_simps = map prover
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   693
    ["(EX x:A. P x & Q) = ((EX x:A. P x) & Q)",
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   694
     "(EX x:A. P & Q x) = (P & (EX x:A. Q x))",
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   695
     "(EX x:{}. P x) = False",
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   696
     "(EX x:insert a B. P x) = (P(a) | (EX x:B. P x))",
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   697
     "(EX x:Union(A). P x) = (EX y:A. EX x:y.  P x)",
3860
a29ab43f7174 More lemmas, esp. ~Bex and ~Ball conversions.
nipkow
parents: 3842
diff changeset
   698
     "(EX x:Collect Q. P x) = (EX x. Q x & P x)",
a29ab43f7174 More lemmas, esp. ~Bex and ~Ball conversions.
nipkow
parents: 3842
diff changeset
   699
     "(EX x:f``A. P x) = (EX x:A. P(f x))",
a29ab43f7174 More lemmas, esp. ~Bex and ~Ball conversions.
nipkow
parents: 3842
diff changeset
   700
     "(~(EX x:A. P x)) = (ALL x:A. ~P x)"];
2513
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   701
3426
9aa5864a7eea The name bex_conj_distrib was WRONG
paulson
parents: 3422
diff changeset
   702
val bex_disj_distrib = 
2513
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   703
    prover "(EX x:A. P x | Q x) = ((EX x:A. P x) | (EX x:A. Q x))";
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   704
2021
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   705
end;
dd5866263153 Added miniscoping for UN and INT
paulson
parents: 1917
diff changeset
   706
2513
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   707
Addsimps (UN1_simps @ INT1_simps @ UN_simps @ INT_simps @ 
d708d8cdc8e8 New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents: 2512
diff changeset
   708
	  ball_simps @ bex_simps);