src/HOL/equalities.ML
author nipkow
Fri, 17 May 1996 16:08:06 +0200
changeset 1748 88650ba93c10
parent 1660 8cb42cd97579
child 1754 852093aeb0ab
permissions -rw-r--r--
Added if_image_distrib.
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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
val eq_cs = set_cs addSIs [equalityI];
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
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
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    15
goal Set.thy "{x.False} = {}";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    16
by (fast_tac eq_cs 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
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    20
goal Set.thy "(A <= {}) = (A = {})";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    21
by (fast_tac eq_cs 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
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    25
section ":";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
goal Set.thy "x ~: {}";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    28
by (fast_tac set_cs 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
qed "in_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    30
Addsimps[in_empty];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
goal Set.thy "x : insert y A = (x=y | x:A)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    33
by (fast_tac set_cs 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
qed "in_insert";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    35
Addsimps[in_insert];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    37
section "insert";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    39
(*NOT SUITABLE FOR REWRITING since {a} == insert a {}*)
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    40
goal Set.thy "insert a A = {a} Un A";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    41
by (fast_tac eq_cs 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    42
qed "insert_is_Un";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    43
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
    44
goal Set.thy "insert a A ~= {}";
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
    45
by (fast_tac (set_cs addEs [equalityCE]) 1);
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
    46
qed"insert_not_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    47
Addsimps[insert_not_empty];
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
    48
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
    49
bind_thm("empty_not_insert",insert_not_empty RS not_sym);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    50
Addsimps[empty_not_insert];
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
    51
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
goal Set.thy "!!a. a:A ==> insert a A = A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
qed "insert_absorb";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    56
goal Set.thy "insert x (insert x A) = insert x A";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    57
by (fast_tac eq_cs 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    58
qed "insert_absorb2";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    59
Addsimps [insert_absorb2];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    60
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
goal Set.thy "(insert x A <= B) = (x:B & A <= B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
by (fast_tac set_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
qed "insert_subset";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    64
Addsimps[insert_subset];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    65
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    66
(* use new B rather than (A-{a}) to avoid infinite unfolding *)
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    67
goal Set.thy "!!a. a:A ==> ? B. A = insert a B & a ~: B";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    68
by (res_inst_tac [("x","A-{a}")] exI 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
    69
by (fast_tac eq_cs 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    70
qed "mk_disjoint_insert";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
    72
section "``";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
goal Set.thy "f``{} = {}";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
qed "image_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    77
Addsimps[image_empty];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
goal Set.thy "f``insert a B = insert (f a) (f``B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
qed "image_insert";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    82
Addsimps[image_insert];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
    84
qed_goal "ball_image" Set.thy "(!y:F``S. P y) = (!x:S. P (F x))"
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
    85
 (fn _ => [fast_tac set_cs 1]);
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
    86
1748
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
    87
goalw Set.thy [image_def]
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
    88
"(%x. if P x then f x else g x) `` S			\
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
    89
\ = (f `` ({x.x:S & P x})) Un (g `` ({x.x:S & ~(P x)}))";
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
    90
by(split_tac [expand_if] 1);
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
    91
by(fast_tac eq_cs 1);
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
    92
qed "if_image_distrib";
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
    93
Addsimps[if_image_distrib];
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
    94
88650ba93c10 Added if_image_distrib.
nipkow
parents: 1660
diff changeset
    95
1660
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
    96
section "range";
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
    97
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
    98
qed_goal "ball_range" Set.thy "(!y:range f. P y) = (!x. P (f x))"
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
    99
 (fn _ => [fast_tac set_cs 1]);
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   100
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   101
qed_goalw "image_range" Set.thy [image_def, range_def]
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   102
 "f``range g = range (%x. f (g x))" (fn _ => [
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   103
	rtac Collect_cong 1,
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   104
	fast_tac set_cs 1]);
8cb42cd97579 *** empty log message ***
oheimb
parents: 1618
diff changeset
   105
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   106
section "Int";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
goal Set.thy "A Int A = A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
qed "Int_absorb";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   111
Addsimps[Int_absorb];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
goal Set.thy "A Int B  =  B Int A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   114
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
qed "Int_commute";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
goal Set.thy "(A Int B) Int C  =  A Int (B Int C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   118
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
qed "Int_assoc";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
goal Set.thy "{} Int B = {}";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
qed "Int_empty_left";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   124
Addsimps[Int_empty_left];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
goal Set.thy "A Int {} = {}";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
qed "Int_empty_right";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   129
Addsimps[Int_empty_right];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   130
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   131
goal Set.thy "UNIV Int B = B";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   132
by (fast_tac eq_cs 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   133
qed "Int_UNIV_left";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   134
Addsimps[Int_UNIV_left];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   135
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   136
goal Set.thy "A Int UNIV = A";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   137
by (fast_tac eq_cs 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   138
qed "Int_UNIV_right";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   139
Addsimps[Int_UNIV_right];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   141
goal Set.thy "A Int (B Un C)  =  (A Int B) Un (A Int C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   142
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   143
qed "Int_Un_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   144
1618
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1564
diff changeset
   145
goal Set.thy "(B Un C) Int A =  (B Int A) Un (C Int A)";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1564
diff changeset
   146
by (fast_tac eq_cs 1);
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1564
diff changeset
   147
qed "Int_Un_distrib2";
372880456b5b Library changes for mutilated checkerboard
paulson
parents: 1564
diff changeset
   148
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   149
goal Set.thy "(A<=B) = (A Int B = A)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   150
by (fast_tac (eq_cs addSEs [equalityE]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   151
qed "subset_Int_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   152
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   153
goal Set.thy "(A Int B = UNIV) = (A = UNIV & B = UNIV)";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   154
by (fast_tac (eq_cs addEs [equalityCE]) 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   155
qed "Int_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   156
Addsimps[Int_UNIV];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   157
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   158
section "Un";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   159
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   160
goal Set.thy "A Un A = A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   161
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   162
qed "Un_absorb";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   163
Addsimps[Un_absorb];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   164
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   165
goal Set.thy "A Un B  =  B Un A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   166
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   167
qed "Un_commute";
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
goal Set.thy "(A Un B) Un C  =  A Un (B Un C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   170
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   171
qed "Un_assoc";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   172
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   173
goal Set.thy "{} Un B = B";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   174
by (fast_tac eq_cs 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   175
qed "Un_empty_left";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   176
Addsimps[Un_empty_left];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
goal Set.thy "A Un {} = A";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   179
by (fast_tac eq_cs 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   180
qed "Un_empty_right";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   181
Addsimps[Un_empty_right];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   182
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   183
goal Set.thy "UNIV Un B = UNIV";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   184
by (fast_tac eq_cs 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   185
qed "Un_UNIV_left";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   186
Addsimps[Un_UNIV_left];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   187
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   188
goal Set.thy "A Un UNIV = UNIV";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   189
by (fast_tac eq_cs 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   190
qed "Un_UNIV_right";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   191
Addsimps[Un_UNIV_right];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   193
goal Set.thy "insert a B Un C = insert a (B Un C)";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   194
by (fast_tac eq_cs 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
qed "Un_insert_left";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   196
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   197
goal Set.thy "(A Int B) Un C  =  (A Un C) Int (B Un C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   199
qed "Un_Int_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   200
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   201
goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   202
 "(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   203
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   204
qed "Un_Int_crazy";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   205
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   206
goal Set.thy "(A<=B) = (A Un B = B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   207
by (fast_tac (eq_cs addSEs [equalityE]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   208
qed "subset_Un_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   209
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   210
goal Set.thy "(A <= insert b C) = (A <= C | b:A & A-{b} <= C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   211
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   212
qed "subset_insert_iff";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   213
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   214
goal Set.thy "(A Un B = {}) = (A = {} & B = {})";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   215
by (fast_tac (eq_cs addEs [equalityCE]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   216
qed "Un_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   217
Addsimps[Un_empty];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   218
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   219
section "Compl";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   220
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   221
goal Set.thy "A Int Compl(A) = {}";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   222
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   223
qed "Compl_disjoint";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   224
Addsimps[Compl_disjoint];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   225
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   226
goal Set.thy "A Un Compl(A) = UNIV";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   227
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   228
qed "Compl_partition";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   229
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   230
goal Set.thy "Compl(Compl(A)) = A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   231
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   232
qed "double_complement";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   233
Addsimps[double_complement];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   234
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   235
goal Set.thy "Compl(A Un B) = Compl(A) Int Compl(B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   236
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   237
qed "Compl_Un";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   238
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   239
goal Set.thy "Compl(A Int B) = Compl(A) Un Compl(B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   240
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   241
qed "Compl_Int";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   242
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   243
goal Set.thy "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   244
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   245
qed "Compl_UN";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   246
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   247
goal Set.thy "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   248
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   249
qed "Compl_INT";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   250
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   251
(*Halmos, Naive Set Theory, page 16.*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   252
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   253
goal Set.thy "((A Int B) Un C = A Int (B Un C)) = (C<=A)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   254
by (fast_tac (eq_cs addSEs [equalityE]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   255
qed "Un_Int_assoc_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   256
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   257
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   258
section "Union";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   259
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   260
goal Set.thy "Union({}) = {}";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   261
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   262
qed "Union_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   263
Addsimps[Union_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   264
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   265
goal Set.thy "Union(UNIV) = UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   266
by (fast_tac eq_cs 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   267
qed "Union_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   268
Addsimps[Union_UNIV];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   269
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   270
goal Set.thy "Union(insert a B) = a Un Union(B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   271
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   272
qed "Union_insert";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   273
Addsimps[Union_insert];
923
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
goal Set.thy "Union(A Un B) = Union(A) Un Union(B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   276
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   277
qed "Union_Un_distrib";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   278
Addsimps[Union_Un_distrib];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   279
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   280
goal Set.thy "Union(A Int B) <= Union(A) Int Union(B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   281
by (fast_tac set_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   282
qed "Union_Int_subset";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   283
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   284
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   285
   "(Union(C) Int A = {}) = (! B:C. B Int A = {})";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   286
by (fast_tac (eq_cs addSEs [equalityE]) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   287
qed "Union_disjoint";
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 "Inter";
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   290
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   291
goal Set.thy "Inter({}) = UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   292
by (fast_tac eq_cs 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   293
qed "Inter_empty";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   294
Addsimps[Inter_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   295
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   296
goal Set.thy "Inter(UNIV) = {}";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   297
by (fast_tac eq_cs 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   298
qed "Inter_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   299
Addsimps[Inter_UNIV];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   300
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   301
goal Set.thy "Inter(insert a B) = a Int Inter(B)";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   302
by (fast_tac eq_cs 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   303
qed "Inter_insert";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   304
Addsimps[Inter_insert];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   305
1564
822575c737bd Deleted faulty comment; proved new rule Inter_Un_subset
paulson
parents: 1553
diff changeset
   306
goal Set.thy "Inter(A) Un Inter(B) <= Inter(A Int B)";
822575c737bd Deleted faulty comment; proved new rule Inter_Un_subset
paulson
parents: 1553
diff changeset
   307
by (fast_tac set_cs 1);
822575c737bd Deleted faulty comment; proved new rule Inter_Un_subset
paulson
parents: 1553
diff changeset
   308
qed "Inter_Un_subset";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   309
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   310
goal Set.thy "Inter(A Un B) = Inter(A) Int Inter(B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   311
by (best_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   312
qed "Inter_Un_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   313
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   314
section "UN and INT";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   315
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   316
(*Basic identities*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   317
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   318
goal Set.thy "(UN x:{}. B x) = {}";
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   319
by (fast_tac eq_cs 1);
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   320
qed "UN_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   321
Addsimps[UN_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   322
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   323
goal Set.thy "(UN x:UNIV. B x) = (UN x. B x)";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   324
by (fast_tac eq_cs 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   325
qed "UN_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   326
Addsimps[UN_UNIV];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   327
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   328
goal Set.thy "(INT x:{}. B x) = UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   329
by (fast_tac eq_cs 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   330
qed "INT_empty";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   331
Addsimps[INT_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   332
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   333
goal Set.thy "(INT x:UNIV. B x) = (INT x. B x)";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   334
by (fast_tac eq_cs 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   335
qed "INT_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   336
Addsimps[INT_UNIV];
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   337
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   338
goal Set.thy "(UN x:insert a A. B x) = B a Un UNION A B";
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   339
by (fast_tac eq_cs 1);
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   340
qed "UN_insert";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   341
Addsimps[UN_insert];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   342
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   343
goal Set.thy "(INT x:insert a A. B x) = B a Int INTER A B";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   344
by (fast_tac eq_cs 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   345
qed "INT_insert";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   346
Addsimps[INT_insert];
1179
7678408f9751 Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents: 923
diff changeset
   347
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   348
goal Set.thy "Union(range(f)) = (UN x.f(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   349
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   350
qed "Union_range_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   351
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   352
goal Set.thy "Inter(range(f)) = (INT x.f(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   353
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   354
qed "Inter_range_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   355
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   356
goal Set.thy "Union(B``A) = (UN x:A. B(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   357
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   358
qed "Union_image_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   359
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   360
goal Set.thy "Inter(B``A) = (INT x:A. B(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   361
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   362
qed "Inter_image_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   363
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   364
goal Set.thy "!!A. a: A ==> (UN y:A. c) = c";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   365
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   366
qed "UN_constant";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   367
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   368
goal Set.thy "!!A. a: A ==> (INT y:A. c) = c";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   369
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   370
qed "INT_constant";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   371
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   372
goal Set.thy "(UN x.B) = B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   373
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   374
qed "UN1_constant";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   375
Addsimps[UN1_constant];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   376
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   377
goal Set.thy "(INT x.B) = B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   378
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   379
qed "INT1_constant";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   380
Addsimps[INT1_constant];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   381
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   382
goal Set.thy "(UN x:A. B(x)) = Union({Y. ? x:A. Y=B(x)})";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   383
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   384
qed "UN_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   385
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   386
(*Look: it has an EXISTENTIAL quantifier*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   387
goal Set.thy "(INT x:A. B(x)) = Inter({Y. ? x:A. Y=B(x)})";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   388
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   389
qed "INT_eq";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   390
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   391
(*Distributive laws...*)
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
goal Set.thy "A Int Union(B) = (UN C:B. A Int C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   394
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   395
qed "Int_Union";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   396
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   397
(* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   398
   Union of a family of unions **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   399
goal Set.thy "(UN x:C. A(x) Un B(x)) = Union(A``C)  Un  Union(B``C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   400
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   401
qed "Un_Union_image";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   402
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   403
(*Equivalent version*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   404
goal Set.thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i))  Un  (UN i:I. B(i))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   405
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   406
qed "UN_Un_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   407
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   408
goal Set.thy "A Un Inter(B) = (INT C:B. A Un C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   409
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   410
qed "Un_Inter";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   411
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   412
goal Set.thy "(INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   413
by (best_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   414
qed "Int_Inter_image";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   415
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   416
(*Equivalent version*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   417
goal Set.thy "(INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   418
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   419
qed "INT_Int_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   420
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   421
(*Halmos, Naive Set Theory, page 35.*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   422
goal Set.thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   423
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   424
qed "Int_UN_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   425
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   426
goal Set.thy "B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   427
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   428
qed "Un_INT_distrib";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   429
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   430
goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   431
    "(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   432
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   433
qed "Int_UN_distrib2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   434
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   435
goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   436
    "(INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   437
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   438
qed "Un_INT_distrib2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   439
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   440
section "-";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   441
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   442
goal Set.thy "A-A = {}";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   443
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   444
qed "Diff_cancel";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   445
Addsimps[Diff_cancel];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   446
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   447
goal Set.thy "{}-A = {}";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   448
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   449
qed "empty_Diff";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   450
Addsimps[empty_Diff];
923
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
goal Set.thy "A-{} = A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   453
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   454
qed "Diff_empty";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   455
Addsimps[Diff_empty];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   456
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   457
goal Set.thy "A-UNIV = {}";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   458
by (fast_tac eq_cs 1);
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   459
qed "Diff_UNIV";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   460
Addsimps[Diff_UNIV];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   461
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   462
goal Set.thy "!!x. x~:A ==> A - insert x B = A-B";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   463
by (fast_tac eq_cs 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   464
qed "Diff_insert0";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   465
Addsimps [Diff_insert0];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   466
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   467
(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   468
goal Set.thy "A - insert a B = A - B - {a}";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   469
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   470
qed "Diff_insert";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   471
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   472
(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   473
goal Set.thy "A - insert a B = A - {a} - B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   474
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   475
qed "Diff_insert2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   476
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   477
goal Set.thy "insert x A - B = (if x:B then A-B else insert x (A-B))";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   478
by (simp_tac (!simpset setloop split_tac[expand_if]) 1);
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   479
by (fast_tac eq_cs 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   480
qed "insert_Diff_if";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   481
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   482
goal Set.thy "!!x. x:B ==> insert x A - B = A-B";
1553
4eb4a9c7d736 Ran expandshort
paulson
parents: 1548
diff changeset
   483
by (fast_tac eq_cs 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   484
qed "insert_Diff1";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   485
Addsimps [insert_Diff1];
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   486
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   487
val prems = goal Set.thy "a:A ==> insert a (A-{a}) = A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   488
by (fast_tac (eq_cs addSIs prems) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   489
qed "insert_Diff";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   490
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   491
goal Set.thy "A Int (B-A) = {}";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   492
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   493
qed "Diff_disjoint";
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   494
Addsimps[Diff_disjoint];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   495
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   496
goal Set.thy "!!A. A<=B ==> A Un (B-A) = B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   497
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   498
qed "Diff_partition";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   499
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   500
goal Set.thy "!!A. [| A<=B; B<= C |] ==> (B - (C - A)) = (A :: 'a set)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   501
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   502
qed "double_diff";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   503
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   504
goal Set.thy "A - (B Un C) = (A-B) Int (A-C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   505
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   506
qed "Diff_Un";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   507
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   508
goal Set.thy "A - (B Int C) = (A-B) Un (A-C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   509
by (fast_tac eq_cs 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   510
qed "Diff_Int";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   511
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   512
Addsimps[subset_UNIV, empty_subsetI, subset_refl];