src/HOL/Subst/Setplus.ML
author clasohm
Wed, 04 Oct 1995 13:12:14 +0100
changeset 1266 3ae9fe3c0f68
parent 968 3cdaa8724175
child 1465 5d7a7e439cec
permissions -rw-r--r--
added local simpsets
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     1
(*  Title: 	Substitutions/setplus.ML
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     2
    Author: 	Martin Coen, Cambridge University Computer Laboratory
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     3
    Copyright   1993  University of Cambridge
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     4
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     5
For setplus.thy.
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     6
Properties of subsets and empty sets.
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     7
*)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     8
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     9
open Setplus;
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    10
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    11
(*********)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    12
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    13
(*** Rules for subsets ***)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    14
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    15
goal Set.thy "A <= B =  (! t.t:A --> t:B)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    16
by (fast_tac set_cs 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    17
qed "subset_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    18
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    19
goalw Setplus.thy [ssubset_def] "A < B = ((A <= B) & ~(A=B))";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    20
by (rtac refl 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    21
qed "ssubset_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    22
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    23
goal Setplus.thy "((A::'a set) <= B) = ((A < B) | (A=B))";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
diff changeset
    24
by (simp_tac (simpset_of "Fun" addsimps [ssubset_iff]) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    25
by (fast_tac set_cs 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    26
qed "subseteq_iff_subset_eq";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    27
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    28
(*Rule in Modus Ponens style*)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    29
goal Setplus.thy "A < B --> c:A --> c:B";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 968
diff changeset
    30
by (simp_tac (simpset_of "Fun" addsimps [ssubset_iff]) 1);
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    31
by (fast_tac set_cs 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    32
qed "ssubsetD";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    33
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    34
(*********)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    35
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    36
goalw Setplus.thy [empty_def] "~ a : {}";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    37
by (fast_tac set_cs 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    38
qed "not_in_empty";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    39
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    40
goalw Setplus.thy [empty_def] "(A = {}) = (ALL a.~ a:A)";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    41
by (fast_tac (set_cs addIs [set_ext]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    42
qed "empty_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    43
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    44
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    45
(*********)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    46
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    47
goal Set.thy "(~A=B)  = ((? x.x:A & ~x:B) | (? x.~x:A & x:B))";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    48
by (fast_tac (set_cs addIs [set_ext]) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    49
qed "not_equal_iff";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    50
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    51
(*********)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    52
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    53
val setplus_rews = [ssubset_iff,not_in_empty,empty_iff];
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    54
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    55
(*********)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    56
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    57
(*Case analysis for rewriting; P also gets rewritten*)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    58
val [prem1,prem2] = goal HOL.thy "[| P-->Q;  ~P-->Q |] ==> Q";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    59
by (rtac (excluded_middle RS disjE) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    60
by (etac (prem2 RS mp) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    61
by (etac (prem1 RS mp) 1);
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    62
qed "imp_excluded_middle";
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    63
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    64
fun imp_excluded_middle_tac s = res_inst_tac [("P",s)] imp_excluded_middle;