src/HOL/subset.ML
author paulson
Mon, 07 Oct 1996 10:28:44 +0200
changeset 2056 93c093620c28
parent 1761 29e08d527ba1
child 2515 6ff9bd353121
permissions -rw-r--r--
Removed commands made redundant by new one-point rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
     1
(*  Title:      HOL/subset
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
Derived rules involving subsets
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
Union and Intersection as lattice operations
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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
(*** insert ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
qed_goal "subset_insertI" Set.thy "B <= insert a B"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
 (fn _=> [ (rtac subsetI 1), (etac insertI2 1) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    15
goal Set.thy "!!x. x ~: A ==> (A <= insert x B) = (A <= B)";
1760
6f41a494f3b1 Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents: 1631
diff changeset
    16
by (Fast_tac 1);
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    17
qed "subset_insert";
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
    18
1631
26570790f089 new lemma for mutilated chess board
paulson
parents: 1552
diff changeset
    19
qed_goal "subset_empty_iff" Set.thy "(A<={}) = (A={})"
1761
29e08d527ba1 Removed equalityI from some proofs (because it is now included
berghofe
parents: 1760
diff changeset
    20
 (fn _=> [ (Fast_tac 1) ]);
1631
26570790f089 new lemma for mutilated chess board
paulson
parents: 1552
diff changeset
    21
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
(*** Big Union -- least upper bound of a set  ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
    "B:A ==> B <= Union(A)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
by (REPEAT (ares_tac (prems@[subsetI,UnionI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
qed "Union_upper";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
val [prem] = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
    "[| !!X. X:A ==> X<=C |] ==> Union(A) <= C";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
    31
by (rtac subsetI 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
by (REPEAT (eresolve_tac [asm_rl, UnionE, prem RS subsetD] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
qed "Union_least";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
(** General union **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
val prems = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
    "a:A ==> B(a) <= (UN x:A. B(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
by (REPEAT (ares_tac (prems@[UN_I RS subsetI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
qed "UN_upper";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
val [prem] = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
    "[| !!x. x:A ==> B(x)<=C |] ==> (UN x:A. B(x)) <= C";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
    44
by (rtac subsetI 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
by (REPEAT (eresolve_tac [asm_rl, UN_E, prem RS subsetD] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
qed "UN_least";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
goal Set.thy "B(a) <= (UN x. B(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
by (REPEAT (ares_tac [UN1_I RS subsetI] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
qed "UN1_upper";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
val [prem] = goal Set.thy "[| !!x. B(x)<=C |] ==> (UN x. B(x)) <= C";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
    53
by (rtac subsetI 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
by (REPEAT (eresolve_tac [asm_rl, UN1_E, prem RS subsetD] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
qed "UN1_least";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
(*** Big Intersection -- greatest lower bound of a set ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
val prems = goal Set.thy "B:A ==> Inter(A) <= B";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
    61
by (rtac subsetI 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
by (REPEAT (resolve_tac prems 1 ORELSE etac InterD 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
qed "Inter_lower";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
val [prem] = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
    "[| !!X. X:A ==> C<=X |] ==> C <= Inter(A)";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
    67
by (rtac (InterI RS subsetI) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
by (REPEAT (eresolve_tac [asm_rl, prem RS subsetD] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
qed "Inter_greatest";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
val prems = goal Set.thy "a:A ==> (INT x:A. B(x)) <= B(a)";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
    72
by (rtac subsetI 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
by (REPEAT (resolve_tac prems 1 ORELSE etac INT_D 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
qed "INT_lower";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
val [prem] = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
    "[| !!x. x:A ==> C<=B(x) |] ==> C <= (INT x:A. B(x))";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
    78
by (rtac (INT_I RS subsetI) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
by (REPEAT (eresolve_tac [asm_rl, prem RS subsetD] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
qed "INT_greatest";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
goal Set.thy "(INT x. B(x)) <= B(a)";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
    83
by (rtac subsetI 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
by (REPEAT (resolve_tac prems 1 ORELSE etac INT1_D 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
qed "INT1_lower";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
val [prem] = goal Set.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
    "[| !!x. C<=B(x) |] ==> C <= (INT x. B(x))";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
    89
by (rtac (INT1_I RS subsetI) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
by (REPEAT (eresolve_tac [asm_rl, prem RS subsetD] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
qed "INT1_greatest";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    92
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
(*** Finite Union -- the least upper bound of 2 sets ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
goal Set.thy "A <= A Un B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
by (REPEAT (ares_tac [subsetI,UnI1] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
qed "Un_upper1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
goal Set.thy "B <= A Un B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
by (REPEAT (ares_tac [subsetI,UnI2] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   101
qed "Un_upper2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
val prems = goal Set.thy "[| A<=C;  B<=C |] ==> A Un B <= C";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
by (cut_facts_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
by (DEPTH_SOLVE (ares_tac [subsetI] 1 
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
          ORELSE eresolve_tac [UnE,subsetD] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
qed "Un_least";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
(*** Finite Intersection -- the greatest lower bound of 2 sets *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
goal Set.thy "A Int B <= A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
by (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
qed "Int_lower1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   114
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
goal Set.thy "A Int B <= B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
by (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
qed "Int_lower2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   118
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
val prems = goal Set.thy "[| C<=A;  C<=B |] ==> C <= A Int B";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
by (cut_facts_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
by (REPEAT (ares_tac [subsetI,IntI] 1
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   122
     ORELSE etac subsetD 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
qed "Int_greatest";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   124
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
(*** Set difference ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   126
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   127
qed_goal "Diff_subset" Set.thy "A-B <= (A::'a set)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   128
 (fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac DiffE 1)) ]);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   129
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   130
(*** Monotonicity ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   131
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   132
val [prem] = goal Set.thy "mono(f) ==> f(A) Un f(B) <= f(A Un B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   133
by (rtac Un_least 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   134
by (rtac (Un_upper1 RS (prem RS monoD)) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   135
by (rtac (Un_upper2 RS (prem RS monoD)) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   136
qed "mono_Un";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   137
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   138
val [prem] = goal Set.thy "mono(f) ==> f(A Int B) <= f(A) Int f(B)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   139
by (rtac Int_greatest 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
by (rtac (Int_lower1 RS (prem RS monoD)) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   141
by (rtac (Int_lower2 RS (prem RS monoD)) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   142
qed "mono_Int";