src/HOL/Set.ML
author nipkow
Tue, 27 Mar 2001 13:00:30 +0200
changeset 11227 d9bda7cbdf56
parent 11223 fef9da0ee890
child 11232 558a4feebb04
permissions -rw-r--r--
fixed bug in tactic for ball 1 point simproc
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9422
4b6bc2b347e5 avoid referencing thy value;
wenzelm
parents: 9378
diff changeset
     1
(*  Title:      HOL/Set.ML
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
1985
84cf16192e03 Tidied many proofs, using AddIffs to let equivalences take
paulson
parents: 1937
diff changeset
     6
Set theory for higher-order logic.  A set is simply a predicate.
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
     9
section "Relating predicates and sets";
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    10
3469
61d927bd57ec Now Collect_mem_eq is a default simprule (how could it have ever been omitted?
paulson
parents: 3420
diff changeset
    11
Addsimps [Collect_mem_eq];
61d927bd57ec Now Collect_mem_eq is a default simprule (how could it have ever been omitted?
paulson
parents: 3420
diff changeset
    12
AddIffs  [mem_Collect_eq];
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    13
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
    14
Goal "P(a) ==> a : {x. P(x)}";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    15
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
qed "CollectI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
    18
Goal "a : {x. P(x)} ==> P(a)";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    19
by (Asm_full_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
qed "CollectD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
9687
772ac061bd76 moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
paulson
parents: 9422
diff changeset
    22
val [prem] = Goal "(!!x. (x:A) = (x:B)) ==> A = B";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
by (rtac (prem RS ext RS arg_cong RS box_equals) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
by (rtac Collect_mem_eq 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
by (rtac Collect_mem_eq 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
qed "set_ext";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
9687
772ac061bd76 moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
paulson
parents: 9422
diff changeset
    28
val [prem] = Goal "(!!x. P(x)=Q(x)) ==> {x. P(x)} = {x. Q(x)}";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
by (rtac (prem RS ext RS arg_cong) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
qed "Collect_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9088
diff changeset
    32
bind_thm ("CollectE", make_elim CollectD);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    34
AddSIs [CollectI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    35
AddSEs [CollectE];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    36
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    37
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
    38
section "Bounded quantifiers";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
    40
val prems = Goalw [Ball_def]
9041
paulson
parents: 8913
diff changeset
    41
    "[| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
by (REPEAT (ares_tac (prems @ [allI,impI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
qed "ballI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
8839
31da5b9790c0 strip = impI allI allI;
wenzelm
parents: 8326
diff changeset
    45
bind_thms ("strip", [impI, allI, ballI]);
31da5b9790c0 strip = impI allI allI;
wenzelm
parents: 8326
diff changeset
    46
9041
paulson
parents: 8913
diff changeset
    47
Goalw [Ball_def] "[| ALL x:A. P(x);  x:A |] ==> P(x)";
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
    48
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
qed "bspec";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
    51
val major::prems = Goalw [Ball_def]
9041
paulson
parents: 8913
diff changeset
    52
    "[| ALL x:A. P(x);  P(x) ==> Q;  x~:A ==> Q |] ==> Q";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
by (rtac (major RS spec RS impCE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
qed "ballE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
9041
paulson
parents: 8913
diff changeset
    57
(*Takes assumptions ALL x:A.P(x) and a:A; creates assumption P(a)*)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
fun ball_tac i = etac ballE i THEN contr_tac (i+1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    60
AddSIs [ballI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    61
AddEs  [ballE];
7441
20b3e2d2fcb6 AddXDs [bspec];
wenzelm
parents: 7031
diff changeset
    62
AddXDs [bspec];
5521
7970832271cc added wrapper for bspec
oheimb
parents: 5490
diff changeset
    63
(* gives better instantiation for bound: *)
11166
eca861fd1eb5 simplified definition of wrapper bspec
oheimb
parents: 11007
diff changeset
    64
claset_ref() := claset() addbefore ("bspec", datac bspec 1);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    65
6006
d2e271b8d651 new rule rev_bexI
paulson
parents: 5931
diff changeset
    66
(*Normally the best argument order: P(x) constrains the choice of x:A*)
9041
paulson
parents: 8913
diff changeset
    67
Goalw [Bex_def] "[| P(x);  x:A |] ==> EX x:A. P(x)";
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
    68
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
qed "bexI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
6006
d2e271b8d651 new rule rev_bexI
paulson
parents: 5931
diff changeset
    71
(*The best argument order when there is only one x:A*)
9041
paulson
parents: 8913
diff changeset
    72
Goalw [Bex_def] "[| x:A;  P(x) |] ==> EX x:A. P(x)";
6006
d2e271b8d651 new rule rev_bexI
paulson
parents: 5931
diff changeset
    73
by (Blast_tac 1);
d2e271b8d651 new rule rev_bexI
paulson
parents: 5931
diff changeset
    74
qed "rev_bexI";
d2e271b8d651 new rule rev_bexI
paulson
parents: 5931
diff changeset
    75
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
    76
val prems = Goal 
9041
paulson
parents: 8913
diff changeset
    77
   "[| ALL x:A. ~P(x) ==> P(a);  a:A |] ==> EX x:A. P(x)";
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
    78
by (rtac classical 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
    79
by (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1))  ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
    80
qed "bexCI";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
    82
val major::prems = Goalw [Bex_def]
9041
paulson
parents: 8913
diff changeset
    83
    "[| EX x:A. P(x);  !!x. [| x:A; P(x) |] ==> Q  |] ==> Q";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
by (rtac (major RS exE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
qed "bexE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    88
AddIs  [bexI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    89
AddSEs [bexE];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
    90
3420
02dc9c5b035f New miniscoping rules ball_triv and bex_triv
paulson
parents: 3222
diff changeset
    91
(*Trival rewrite rule*)
9041
paulson
parents: 8913
diff changeset
    92
Goal "(ALL x:A. P) = ((EX x. x:A) --> P)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
    93
by (simp_tac (simpset() addsimps [Ball_def]) 1);
3420
02dc9c5b035f New miniscoping rules ball_triv and bex_triv
paulson
parents: 3222
diff changeset
    94
qed "ball_triv";
1816
b03dba9116d4 New rewrites for vacuous quantification
paulson
parents: 1776
diff changeset
    95
1882
67f49e8c4355 Proved bex_False
paulson
parents: 1841
diff changeset
    96
(*Dual form for existentials*)
9041
paulson
parents: 8913
diff changeset
    97
Goal "(EX x:A. P) = ((EX x. x:A) & P)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
    98
by (simp_tac (simpset() addsimps [Bex_def]) 1);
3420
02dc9c5b035f New miniscoping rules ball_triv and bex_triv
paulson
parents: 3222
diff changeset
    99
qed "bex_triv";
1882
67f49e8c4355 Proved bex_False
paulson
parents: 1841
diff changeset
   100
3420
02dc9c5b035f New miniscoping rules ball_triv and bex_triv
paulson
parents: 3222
diff changeset
   101
Addsimps [ball_triv, bex_triv];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
11223
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   103
Goal "(EX x:A. x=a) = (a:A)";
11220
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   104
by(Blast_tac 1);
11223
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   105
qed "bex_triv_one_point1";
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   106
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   107
Goal "(EX x:A. a=x) = (a:A)";
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   108
by(Blast_tac 1);
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   109
qed "bex_triv_one_point2";
11220
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   110
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   111
Goal "(EX x:A. x=a & P x) = (a:A & P a)";
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   112
by(Blast_tac 1);
11223
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   113
qed "bex_one_point1";
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   114
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   115
Goal "(ALL x:A. x=a --> P x) = (a:A --> P a)";
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   116
by(Blast_tac 1);
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   117
qed "ball_one_point1";
11220
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   118
11223
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   119
Goal "(ALL x:A. a=x --> P x) = (a:A --> P a)";
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   120
by(Blast_tac 1);
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   121
qed "ball_one_point2";
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   122
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   123
Addsimps [bex_triv_one_point1,bex_triv_one_point2,bex_one_point1,
fef9da0ee890 I forgot a few bases cases for the 1-point rules...
nipkow
parents: 11220
diff changeset
   124
          ball_one_point1,ball_one_point2];
11220
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   125
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   126
let
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   127
val ex_pattern = Thm.read_cterm (Theory.sign_of (the_context ()))
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   128
    ("EX x:A. P(x) & Q(x)",HOLogic.boolT)
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   129
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   130
val prove_bex_tac = rewrite_goals_tac [Bex_def] THEN
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   131
                    Quantifier1.prove_one_point_ex_tac;
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   132
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   133
val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac;
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   134
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   135
val all_pattern = Thm.read_cterm (Theory.sign_of (the_context ()))
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   136
    ("ALL x:A. P(x) & P'(x) --> Q(x)",HOLogic.boolT)
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   137
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   138
val swap = prove_goal thy
11227
d9bda7cbdf56 fixed bug in tactic for ball 1 point simproc
nipkow
parents: 11223
diff changeset
   139
  "((!x. A x & EP x --> Q x) = (!x. E x --> A x & P x --> Q x)) ==> \
d9bda7cbdf56 fixed bug in tactic for ball 1 point simproc
nipkow
parents: 11223
diff changeset
   140
\  ((!x. A x --> EP x --> Q x) = (!x. A x --> E x --> P x --> Q x))"
11220
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   141
  (fn ths => [cut_facts_tac ths 1, Blast_tac 1]);
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   142
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   143
val prove_ball_tac = rewrite_goals_tac [Ball_def] THEN rtac swap 1 THEN
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   144
                     Quantifier1.prove_one_point_all_tac;
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   145
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   146
val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac;
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   147
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   148
val defBEX_regroup = mk_simproc "defined BEX" [ex_pattern] rearrange_bex;
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   149
val defBALL_regroup = mk_simproc "defined BALL" [all_pattern] rearrange_ball;
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   150
in
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   151
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   152
Addsimprocs [defBALL_regroup,defBEX_regroup]
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   153
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   154
end;
db536a42dfc5 added one point simprocs for bounded quantifiers
nipkow
parents: 11166
diff changeset
   155
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   156
(** Congruence rules **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   157
6291
2c3f72d9f5d1 simpler proofs of congruence rules
paulson
parents: 6171
diff changeset
   158
val prems = Goalw [Ball_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   159
    "[| A=B;  !!x. x:B ==> P(x) = Q(x) |] ==> \
9041
paulson
parents: 8913
diff changeset
   160
\    (ALL x:A. P(x)) = (ALL x:B. Q(x))";
6291
2c3f72d9f5d1 simpler proofs of congruence rules
paulson
parents: 6171
diff changeset
   161
by (asm_simp_tac (simpset() addsimps prems) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   162
qed "ball_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   163
6291
2c3f72d9f5d1 simpler proofs of congruence rules
paulson
parents: 6171
diff changeset
   164
val prems = Goalw [Bex_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   165
    "[| A=B;  !!x. x:B ==> P(x) = Q(x) |] ==> \
9041
paulson
parents: 8913
diff changeset
   166
\    (EX x:A. P(x)) = (EX x:B. Q(x))";
6291
2c3f72d9f5d1 simpler proofs of congruence rules
paulson
parents: 6171
diff changeset
   167
by (asm_simp_tac (simpset() addcongs [conj_cong] addsimps prems) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   168
qed "bex_cong";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   169
6291
2c3f72d9f5d1 simpler proofs of congruence rules
paulson
parents: 6171
diff changeset
   170
Addcongs [ball_cong,bex_cong];
2c3f72d9f5d1 simpler proofs of congruence rules
paulson
parents: 6171
diff changeset
   171
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   172
section "Subsets";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   173
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   174
val prems = Goalw [subset_def] "(!!x. x:A ==> x:B) ==> A <= B";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   175
by (REPEAT (ares_tac (prems @ [ballI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   176
qed "subsetI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
5649
1bac26652f45 Uses overload_1st_set to specify overloading
paulson
parents: 5600
diff changeset
   178
(*Map the type ('a set => anything) to just 'a.
1bac26652f45 Uses overload_1st_set to specify overloading
paulson
parents: 5600
diff changeset
   179
  For overloading constants whose first argument has type "'a set" *)
1bac26652f45 Uses overload_1st_set to specify overloading
paulson
parents: 5600
diff changeset
   180
fun overload_1st_set s = Blast.overloaded (s, HOLogic.dest_setT o domain_type);
1bac26652f45 Uses overload_1st_set to specify overloading
paulson
parents: 5600
diff changeset
   181
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   182
(*While (:) is not, its type must be kept
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   183
  for overloading of = to work.*)
4240
8ba60a4cd380 Renamed "overload" to "overloaded" for sml/nj compatibility
paulson
parents: 4231
diff changeset
   184
Blast.overloaded ("op :", domain_type);
5649
1bac26652f45 Uses overload_1st_set to specify overloading
paulson
parents: 5600
diff changeset
   185
1bac26652f45 Uses overload_1st_set to specify overloading
paulson
parents: 5600
diff changeset
   186
overload_1st_set "Ball";		(*need UNION, INTER also?*)
1bac26652f45 Uses overload_1st_set to specify overloading
paulson
parents: 5600
diff changeset
   187
overload_1st_set "Bex";
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   188
4469
399868bf8444 Overloading info for image
paulson
parents: 4434
diff changeset
   189
(*Image: retain the type of the set being expressed*)
8005
b64d86018785 new-style infix declaration for "image"
paulson
parents: 8001
diff changeset
   190
Blast.overloaded ("image", domain_type);
2881
62ecde1015ae Declares overloading for set-theoretic constants
paulson
parents: 2858
diff changeset
   191
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
(*Rule in Modus Ponens style*)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   193
Goalw [subset_def] "[| A <= B;  c:A |] ==> c:B";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   194
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
qed "subsetD";
7658
2d3445be4e91 CollectE;
wenzelm
parents: 7499
diff changeset
   196
AddXIs [subsetD];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   197
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
(*The same, with reversed premises for use with etac -- cf rev_mp*)
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   199
Goal "[| c:A;  A <= B |] ==> c:B";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   200
by (REPEAT (ares_tac [subsetD] 1)) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   201
qed "rev_subsetD";
7658
2d3445be4e91 CollectE;
wenzelm
parents: 7499
diff changeset
   202
AddXIs [rev_subsetD];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   203
1920
df683ce7aad8 Added impOfSubs
paulson
parents: 1882
diff changeset
   204
(*Converts A<=B to x:A ==> x:B*)
df683ce7aad8 Added impOfSubs
paulson
parents: 1882
diff changeset
   205
fun impOfSubs th = th RSN (2, rev_subsetD);
df683ce7aad8 Added impOfSubs
paulson
parents: 1882
diff changeset
   206
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   207
(*Classical elimination rule*)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   208
val major::prems = Goalw [subset_def] 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   209
    "[| A <= B;  c~:A ==> P;  c:B ==> P |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   210
by (rtac (major RS ballE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   211
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   212
qed "subsetCE";
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
(*Takes assumptions A<=B; c:A and creates the assumption c:B *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   215
fun set_mp_tac i = etac subsetCE i  THEN  mp_tac i;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   216
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   217
AddSIs [subsetI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   218
AddEs  [subsetD, subsetCE];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   219
10233
08083bd2a64d tidying; removed unused rev_contra_subsetD
paulson
parents: 9969
diff changeset
   220
Goal "[| A <= B; c ~: B |] ==> c ~: A";
08083bd2a64d tidying; removed unused rev_contra_subsetD
paulson
parents: 9969
diff changeset
   221
by (Blast_tac 1);
08083bd2a64d tidying; removed unused rev_contra_subsetD
paulson
parents: 9969
diff changeset
   222
qed "contra_subsetD";
08083bd2a64d tidying; removed unused rev_contra_subsetD
paulson
parents: 9969
diff changeset
   223
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   224
Goal "A <= (A::'a set)";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   225
by (Fast_tac 1);
10233
08083bd2a64d tidying; removed unused rev_contra_subsetD
paulson
parents: 9969
diff changeset
   226
qed "subset_refl";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   227
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   228
Goal "[| A<=B;  B<=C |] ==> A<=(C::'a set)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   229
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   230
qed "subset_trans";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   231
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   232
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   233
section "Equality";
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
(*Anti-symmetry of the subset relation*)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   236
Goal "[| A <= B;  B <= A |] ==> A = (B::'a set)";
5318
72bf8039b53f expandshort
paulson
parents: 5316
diff changeset
   237
by (rtac set_ext 1);
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   238
by (blast_tac (claset() addIs [subsetD]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   239
qed "subset_antisym";
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9088
diff changeset
   240
bind_thm ("equalityI", subset_antisym);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   241
1762
6e481897a811 equalityI is now added to default claset
berghofe
parents: 1760
diff changeset
   242
AddSIs [equalityI];
6e481897a811 equalityI is now added to default claset
berghofe
parents: 1760
diff changeset
   243
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   244
(* Equality rules from ZF set theory -- are they appropriate here? *)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   245
Goal "A = B ==> A<=(B::'a set)";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   246
by (etac ssubst 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   247
by (rtac subset_refl 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   248
qed "equalityD1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   249
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   250
Goal "A = B ==> B<=(A::'a set)";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   251
by (etac ssubst 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   252
by (rtac subset_refl 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   253
qed "equalityD2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   254
9338
fcf7f29a3447 re-added subset_empty to simpset
oheimb
parents: 9186
diff changeset
   255
(*Be careful when adding this to the claset as subset_empty is in the simpset:
fcf7f29a3447 re-added subset_empty to simpset
oheimb
parents: 9186
diff changeset
   256
  A={} goes to {}<=A and A<={} and then back to A={} !*)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   257
val prems = Goal
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   258
    "[| A = B;  [| A<=B; B<=(A::'a set) |] ==> P |]  ==>  P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   259
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   260
by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   261
qed "equalityE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   262
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   263
val major::prems = Goal
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   264
    "[| A = B;  [| c:A; c:B |] ==> P;  [| c~:A; c~:B |] ==> P |]  ==>  P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   265
by (rtac (major RS equalityE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   266
by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   267
qed "equalityCE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   268
9186
7b2f4e6538b4 now uses equalityCE, which usually is more efficent than equalityE
paulson
parents: 9161
diff changeset
   269
AddEs [equalityCE];
7b2f4e6538b4 now uses equalityCE, which usually is more efficent than equalityE
paulson
parents: 9161
diff changeset
   270
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   271
(*Lemma for creating induction formulae -- for "pattern matching" on p
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   272
  To make the induction hypotheses usable, apply "spec" or "bspec" to
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   273
  put universal quantifiers over the free variables in p. *)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   274
val prems = Goal 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   275
    "[| p:A;  !!z. z:A ==> p=z --> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   276
by (rtac mp 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   277
by (REPEAT (resolve_tac (refl::prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   278
qed "setup_induction";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   279
8053
37ebdaf3bb91 useful lemma eqset_imp_iff
paulson
parents: 8025
diff changeset
   280
Goal "A = B ==> (x : A) = (x : B)";
37ebdaf3bb91 useful lemma eqset_imp_iff
paulson
parents: 8025
diff changeset
   281
by (Asm_simp_tac 1);
37ebdaf3bb91 useful lemma eqset_imp_iff
paulson
parents: 8025
diff changeset
   282
qed "eqset_imp_iff";
37ebdaf3bb91 useful lemma eqset_imp_iff
paulson
parents: 8025
diff changeset
   283
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   284
4159
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   285
section "The universal set -- UNIV";
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   286
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   287
Goalw [UNIV_def] "x : UNIV";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   288
by (rtac CollectI 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   289
by (rtac TrueI 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   290
qed "UNIV_I";
4159
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   291
4434
75f38104ff80 UNIV_I no longer counts as safe
paulson
parents: 4423
diff changeset
   292
Addsimps [UNIV_I];
75f38104ff80 UNIV_I no longer counts as safe
paulson
parents: 4423
diff changeset
   293
AddIs    [UNIV_I];  (*unsafe makes it less likely to cause problems*)
4159
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   294
10482
41de88cb2108 UNIV_witness;
wenzelm
parents: 10233
diff changeset
   295
Goal "EX x. x : UNIV";
41de88cb2108 UNIV_witness;
wenzelm
parents: 10233
diff changeset
   296
by (Simp_tac 1);
41de88cb2108 UNIV_witness;
wenzelm
parents: 10233
diff changeset
   297
qed "UNIV_witness";
41de88cb2108 UNIV_witness;
wenzelm
parents: 10233
diff changeset
   298
AddXIs [UNIV_witness];
41de88cb2108 UNIV_witness;
wenzelm
parents: 10233
diff changeset
   299
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   300
Goal "A <= UNIV";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   301
by (rtac subsetI 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   302
by (rtac UNIV_I 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   303
qed "subset_UNIV";
4159
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   304
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   305
(** Eta-contracting these two rules (to remove P) causes them to be ignored
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   306
    because of their interaction with congruence rules. **)
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   307
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   308
Goalw [Ball_def] "Ball UNIV P = All P";
4159
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   309
by (Simp_tac 1);
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   310
qed "ball_UNIV";
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   311
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   312
Goalw [Bex_def] "Bex UNIV P = Ex P";
4159
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   313
by (Simp_tac 1);
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   314
qed "bex_UNIV";
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   315
Addsimps [ball_UNIV, bex_UNIV];
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   316
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   317
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   318
section "The empty set -- {}";
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   319
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   320
Goalw [empty_def] "(c : {}) = False";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   321
by (Blast_tac 1) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   322
qed "empty_iff";
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   323
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   324
Addsimps [empty_iff];
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   325
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   326
Goal "a:{} ==> P";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   327
by (Full_simp_tac 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   328
qed "emptyE";
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   329
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   330
AddSEs [emptyE];
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   331
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   332
Goal "{} <= A";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   333
by (Blast_tac 1) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   334
qed "empty_subsetI";
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   335
5256
e6983301ce5e Renamed equals0D to equals0E
paulson
parents: 5237
diff changeset
   336
(*One effect is to delete the ASSUMPTION {} <= A*)
e6983301ce5e Renamed equals0D to equals0E
paulson
parents: 5237
diff changeset
   337
AddIffs [empty_subsetI];
e6983301ce5e Renamed equals0D to equals0E
paulson
parents: 5237
diff changeset
   338
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   339
val [prem]= Goal "[| !!y. y:A ==> False |] ==> A={}";
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   340
by (blast_tac (claset() addIs [prem RS FalseE]) 1) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   341
qed "equals0I";
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   342
5256
e6983301ce5e Renamed equals0D to equals0E
paulson
parents: 5237
diff changeset
   343
(*Use for reasoning about disjointness: A Int B = {} *)
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   344
Goal "A={} ==> a ~: A";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   345
by (Blast_tac 1) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   346
qed "equals0D";
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   347
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   348
Goalw [Ball_def] "Ball {} P = True";
4159
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   349
by (Simp_tac 1);
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   350
qed "ball_empty";
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   351
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   352
Goalw [Bex_def] "Bex {} P = False";
4159
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   353
by (Simp_tac 1);
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   354
qed "bex_empty";
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   355
Addsimps [ball_empty, bex_empty];
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   356
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   357
Goal "UNIV ~= {}";
4159
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   358
by (blast_tac (claset() addEs [equalityE]) 1);
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   359
qed "UNIV_not_empty";
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   360
AddIffs [UNIV_not_empty];
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   361
4aff9b7e5597 UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents: 4135
diff changeset
   362
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   363
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   364
section "The Powerset operator -- Pow";
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   365
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   366
Goalw [Pow_def] "(A : Pow(B)) = (A <= B)";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   367
by (Asm_simp_tac 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   368
qed "Pow_iff";
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   369
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   370
AddIffs [Pow_iff]; 
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   371
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   372
Goalw [Pow_def] "A <= B ==> A : Pow(B)";
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   373
by (etac CollectI 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   374
qed "PowI";
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   375
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   376
Goalw [Pow_def] "A : Pow(B)  ==>  A<=B";
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   377
by (etac CollectD 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   378
qed "PowD";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   379
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   380
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9088
diff changeset
   381
bind_thm ("Pow_bottom", empty_subsetI RS PowI);        (* {}: Pow(B) *)
9fff97d29837 bind_thm(s);
wenzelm
parents: 9088
diff changeset
   382
bind_thm ("Pow_top", subset_refl RS PowI);             (* A : Pow(A) *)
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   383
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   384
5931
325300576da7 Finally removing "Compl" from HOL
paulson
parents: 5649
diff changeset
   385
section "Set complement";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   386
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   387
Goalw [Compl_def] "(c : -A) = (c~:A)";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   388
by (Blast_tac 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   389
qed "Compl_iff";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   390
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   391
Addsimps [Compl_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   392
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5450
diff changeset
   393
val prems = Goalw [Compl_def] "[| c:A ==> False |] ==> c : -A";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   394
by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   395
qed "ComplI";
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
(*This form, with negated conclusion, works well with the Classical prover.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   398
  Negated assumptions behave like formulae on the right side of the notional
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   399
  turnstile...*)
5490
85855f65d0c6 From Compl(A) to -A
paulson
parents: 5450
diff changeset
   400
Goalw [Compl_def] "c : -A ==> c~:A";
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   401
by (etac CollectD 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   402
qed "ComplD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   403
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9088
diff changeset
   404
bind_thm ("ComplE", make_elim ComplD);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   405
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   406
AddSIs [ComplI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   407
AddSEs [ComplE];
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   408
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   409
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   410
section "Binary union -- Un";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   411
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   412
Goalw [Un_def] "(c : A Un B) = (c:A | c:B)";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   413
by (Blast_tac 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   414
qed "Un_iff";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   415
Addsimps [Un_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   416
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   417
Goal "c:A ==> c : A Un B";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   418
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   419
qed "UnI1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   420
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   421
Goal "c:B ==> c : A Un B";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   422
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   423
qed "UnI2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   424
9378
12f251a5a3b5 AddXIs [UnI1, UnI2];
wenzelm
parents: 9338
diff changeset
   425
AddXIs [UnI1, UnI2];
12f251a5a3b5 AddXIs [UnI1, UnI2];
wenzelm
parents: 9338
diff changeset
   426
12f251a5a3b5 AddXIs [UnI1, UnI2];
wenzelm
parents: 9338
diff changeset
   427
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   428
(*Classical introduction rule: no commitment to A vs B*)
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   429
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   430
val prems = Goal "(c~:B ==> c:A) ==> c : A Un B";
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   431
by (Simp_tac 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   432
by (REPEAT (ares_tac (prems@[disjCI]) 1)) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   433
qed "UnCI";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   434
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   435
val major::prems = Goalw [Un_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   436
    "[| c : A Un B;  c:A ==> P;  c:B ==> P |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   437
by (rtac (major RS CollectD RS disjE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   438
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   439
qed "UnE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   440
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   441
AddSIs [UnCI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   442
AddSEs [UnE];
1640
581165679095 Added more _iff rewrites for Compl, Un, Int
paulson
parents: 1618
diff changeset
   443
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   444
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   445
section "Binary intersection -- Int";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   446
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   447
Goalw [Int_def] "(c : A Int B) = (c:A & c:B)";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   448
by (Blast_tac 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   449
qed "Int_iff";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   450
Addsimps [Int_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   451
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   452
Goal "[| c:A;  c:B |] ==> c : A Int B";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   453
by (Asm_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   454
qed "IntI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   455
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   456
Goal "c : A Int B ==> c:A";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   457
by (Asm_full_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   458
qed "IntD1";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   459
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   460
Goal "c : A Int B ==> c:B";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   461
by (Asm_full_simp_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   462
qed "IntD2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   463
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   464
val [major,minor] = Goal
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   465
    "[| c : A Int B;  [| c:A; c:B |] ==> P |] ==> P";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   466
by (rtac minor 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   467
by (rtac (major RS IntD1) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   468
by (rtac (major RS IntD2) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   469
qed "IntE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   470
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   471
AddSIs [IntI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   472
AddSEs [IntE];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   473
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   474
section "Set difference";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   475
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   476
Goalw [set_diff_def] "(c : A-B) = (c:A & c~:B)";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   477
by (Blast_tac 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   478
qed "Diff_iff";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   479
Addsimps [Diff_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   480
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   481
Goal "[| c : A;  c ~: B |] ==> c : A - B";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   482
by (Asm_simp_tac 1) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   483
qed "DiffI";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   484
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   485
Goal "c : A - B ==> c : A";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   486
by (Asm_full_simp_tac 1) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   487
qed "DiffD1";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   488
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   489
Goal "[| c : A - B;  c : B |] ==> P";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   490
by (Asm_full_simp_tac 1) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   491
qed "DiffD2";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   492
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   493
val prems = Goal "[| c : A - B;  [| c:A; c~:B |] ==> P |] ==> P";
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   494
by (resolve_tac prems 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   495
by (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   496
qed "DiffE";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   497
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   498
AddSIs [DiffI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   499
AddSEs [DiffE];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   500
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   501
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   502
section "Augmenting a set -- insert";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   503
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   504
Goalw [insert_def] "a : insert b A = (a=b | a:A)";
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   505
by (Blast_tac 1);
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   506
qed "insert_iff";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   507
Addsimps [insert_iff];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   508
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   509
Goal "a : insert a B";
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   510
by (Simp_tac 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   511
qed "insertI1";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   512
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   513
Goal "!!a. a : B ==> a : insert b B";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   514
by (Asm_simp_tac 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   515
qed "insertI2";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   516
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   517
val major::prems = Goalw [insert_def]
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   518
    "[| a : insert b A;  a=b ==> P;  a:A ==> P |] ==> P";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   519
by (rtac (major RS UnE) 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   520
by (REPEAT (eresolve_tac (prems @ [CollectE]) 1));
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   521
qed "insertE";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   522
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   523
(*Classical introduction rule*)
7031
972b5f62f476 getting rid of qed_goal
paulson
parents: 7007
diff changeset
   524
val prems = Goal "(a~:B ==> a=b) ==> a: insert b B";
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   525
by (Simp_tac 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   526
by (REPEAT (ares_tac (prems@[disjCI]) 1)) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   527
qed "insertCI";
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   528
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   529
AddSIs [insertCI]; 
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   530
AddSEs [insertE];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   531
9088
453996655ac2 replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents: 9075
diff changeset
   532
Goal "(A <= insert x B) = (if x:A then A-{x} <= B else A<=B)";
453996655ac2 replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents: 9075
diff changeset
   533
by Auto_tac; 
453996655ac2 replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents: 9075
diff changeset
   534
qed "subset_insert_iff";
7496
93ae11d887ff added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents: 7441
diff changeset
   535
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   536
section "Singletons, using insert";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   537
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   538
Goal "a : {a}";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   539
by (rtac insertI1 1) ;
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   540
qed "singletonI";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   541
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   542
Goal "b : {a} ==> b=a";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   543
by (Blast_tac 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   544
qed "singletonD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   545
1776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   546
bind_thm ("singletonE", make_elim singletonD);
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   547
7007
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   548
Goal "(b : {a}) = (b=a)";
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   549
by (Blast_tac 1);
b46ccfee8e59 qed_goal -> Goal
paulson
parents: 6443
diff changeset
   550
qed "singleton_iff";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   551
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   552
Goal "{a}={b} ==> a=b";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   553
by (blast_tac (claset() addEs [equalityE]) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   554
qed "singleton_inject";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   555
2858
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   556
(*Redundant? But unlike insertCI, it proves the subgoal immediately!*)
1f3f5c44e159 Re-ordering of rules to assist blast_tac
paulson
parents: 2721
diff changeset
   557
AddSIs [singletonI];   
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   558
AddSDs [singleton_inject];
3718
d78cf498a88c Minor tidying to use Clarify_tac, etc.
paulson
parents: 3582
diff changeset
   559
AddSEs [singletonE];
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   560
7969
7a20317850ab improved singleton_insert_inj_eq
oheimb
parents: 7717
diff changeset
   561
Goal "{b} = insert a A = (a = b & A <= {b})";
8326
0e329578b0ef tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
paulson
parents: 8053
diff changeset
   562
by (blast_tac (claset() addSEs [equalityE]) 1);
7496
93ae11d887ff added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents: 7441
diff changeset
   563
qed "singleton_insert_inj_eq";
93ae11d887ff added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents: 7441
diff changeset
   564
8326
0e329578b0ef tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
paulson
parents: 8053
diff changeset
   565
Goal "(insert a A = {b}) = (a = b & A <= {b})";
0e329578b0ef tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
paulson
parents: 8053
diff changeset
   566
by (blast_tac (claset() addSEs [equalityE]) 1);
7969
7a20317850ab improved singleton_insert_inj_eq
oheimb
parents: 7717
diff changeset
   567
qed "singleton_insert_inj_eq'";
7a20317850ab improved singleton_insert_inj_eq
oheimb
parents: 7717
diff changeset
   568
8326
0e329578b0ef tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
paulson
parents: 8053
diff changeset
   569
AddIffs [singleton_insert_inj_eq, singleton_insert_inj_eq'];
0e329578b0ef tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
paulson
parents: 8053
diff changeset
   570
7496
93ae11d887ff added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents: 7441
diff changeset
   571
Goal "A <= {x} ==> A={} | A = {x}";
93ae11d887ff added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents: 7441
diff changeset
   572
by (Fast_tac 1);
93ae11d887ff added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents: 7441
diff changeset
   573
qed "subset_singletonD";
93ae11d887ff added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents: 7441
diff changeset
   574
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   575
Goal "{x. x=a} = {a}";
4423
a129b817b58a expandshort;
wenzelm
parents: 4240
diff changeset
   576
by (Blast_tac 1);
3582
b87c86b6c291 Added {x.x=a} = a to !simpset.
nipkow
parents: 3469
diff changeset
   577
qed "singleton_conv";
b87c86b6c291 Added {x.x=a} = a to !simpset.
nipkow
parents: 3469
diff changeset
   578
Addsimps [singleton_conv];
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   579
5600
34b3366b83ac new singleton_conv2
nipkow
parents: 5521
diff changeset
   580
Goal "{x. a=x} = {a}";
6301
08245f5a436d expandshort
paulson
parents: 6291
diff changeset
   581
by (Blast_tac 1);
5600
34b3366b83ac new singleton_conv2
nipkow
parents: 5521
diff changeset
   582
qed "singleton_conv2";
34b3366b83ac new singleton_conv2
nipkow
parents: 5521
diff changeset
   583
Addsimps [singleton_conv2];
34b3366b83ac new singleton_conv2
nipkow
parents: 5521
diff changeset
   584
11007
438f31613093 added diff_single_insert and subset_image_iff
oheimb
parents: 10832
diff changeset
   585
Goal "[| A - {x} <= B; x : A |] ==> A <= insert x B"; 
438f31613093 added diff_single_insert and subset_image_iff
oheimb
parents: 10832
diff changeset
   586
by(Blast_tac 1);
438f31613093 added diff_single_insert and subset_image_iff
oheimb
parents: 10832
diff changeset
   587
qed "diff_single_insert";
438f31613093 added diff_single_insert and subset_image_iff
oheimb
parents: 10832
diff changeset
   588
1531
e5eb247ad13c Added a constant UNIV == {x.True}
nipkow
parents: 1465
diff changeset
   589
10832
e33b47e4246d `` -> and ``` -> ``
nipkow
parents: 10482
diff changeset
   590
section "Unions of families -- UNION x:A. B(x) is Union(B`A)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   591
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   592
Goalw [UNION_def] "(b: (UN x:A. B(x))) = (EX x:A. b: B(x))";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   593
by (Blast_tac 1);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   594
qed "UN_iff";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   595
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   596
Addsimps [UN_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   597
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   598
(*The order of the premises presupposes that A is rigid; b may be flexible*)
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   599
Goal "[| a:A;  b: B(a) |] ==> b: (UN x:A. B(x))";
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4469
diff changeset
   600
by Auto_tac;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   601
qed "UN_I";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   602
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   603
val major::prems = Goalw [UNION_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   604
    "[| b : (UN x:A. B(x));  !!x.[| x:A;  b: B(x) |] ==> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   605
by (rtac (major RS CollectD RS bexE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   606
by (REPEAT (ares_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   607
qed "UN_E";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   608
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   609
AddIs  [UN_I];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   610
AddSEs [UN_E];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   611
6291
2c3f72d9f5d1 simpler proofs of congruence rules
paulson
parents: 6171
diff changeset
   612
val prems = Goalw [UNION_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   613
    "[| A=B;  !!x. x:B ==> C(x) = D(x) |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   614
\    (UN x:A. C(x)) = (UN x:B. D(x))";
6291
2c3f72d9f5d1 simpler proofs of congruence rules
paulson
parents: 6171
diff changeset
   615
by (asm_simp_tac (simpset() addsimps prems) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   616
qed "UN_cong";
9687
772ac061bd76 moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
paulson
parents: 9422
diff changeset
   617
Addcongs [UN_cong];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   618
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   619
10832
e33b47e4246d `` -> and ``` -> ``
nipkow
parents: 10482
diff changeset
   620
section "Intersections of families -- INTER x:A. B(x) is Inter(B`A)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   621
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   622
Goalw [INTER_def] "(b: (INT x:A. B(x))) = (ALL x:A. b: B(x))";
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4469
diff changeset
   623
by Auto_tac;
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   624
qed "INT_iff";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   625
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   626
Addsimps [INT_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   627
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   628
val prems = Goalw [INTER_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   629
    "(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   630
by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   631
qed "INT_I";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   632
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   633
Goal "[| b : (INT x:A. B(x));  a:A |] ==> b: B(a)";
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4469
diff changeset
   634
by Auto_tac;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   635
qed "INT_D";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   636
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   637
(*"Classical" elimination -- by the Excluded Middle on a:A *)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   638
val major::prems = Goalw [INTER_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   639
    "[| b : (INT x:A. B(x));  b: B(a) ==> R;  a~:A ==> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   640
by (rtac (major RS CollectD RS ballE) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   641
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   642
qed "INT_E";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   643
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   644
AddSIs [INT_I];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   645
AddEs  [INT_D, INT_E];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   646
6291
2c3f72d9f5d1 simpler proofs of congruence rules
paulson
parents: 6171
diff changeset
   647
val prems = Goalw [INTER_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   648
    "[| A=B;  !!x. x:B ==> C(x) = D(x) |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   649
\    (INT x:A. C(x)) = (INT x:B. D(x))";
6291
2c3f72d9f5d1 simpler proofs of congruence rules
paulson
parents: 6171
diff changeset
   650
by (asm_simp_tac (simpset() addsimps prems) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   651
qed "INT_cong";
9687
772ac061bd76 moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
paulson
parents: 9422
diff changeset
   652
Addcongs [INT_cong];
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   653
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   654
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   655
section "Union";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   656
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   657
Goalw [Union_def] "(A : Union(C)) = (EX X:C. A:X)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   658
by (Blast_tac 1);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   659
qed "Union_iff";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   660
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   661
Addsimps [Union_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   662
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   663
(*The order of the premises presupposes that C is rigid; A may be flexible*)
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   664
Goal "[| X:C;  A:X |] ==> A : Union(C)";
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4469
diff changeset
   665
by Auto_tac;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   666
qed "UnionI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   667
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   668
val major::prems = Goalw [Union_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   669
    "[| A : Union(C);  !!X.[| A:X;  X:C |] ==> R |] ==> R";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   670
by (rtac (major RS UN_E) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   671
by (REPEAT (ares_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   672
qed "UnionE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   673
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   674
AddIs  [UnionI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   675
AddSEs [UnionE];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   676
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   677
1548
afe750876848 Added 'section' commands
nipkow
parents: 1531
diff changeset
   678
section "Inter";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   679
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   680
Goalw [Inter_def] "(A : Inter(C)) = (ALL X:C. A:X)";
2891
d8f254ad1ab9 Calls Blast_tac
paulson
parents: 2881
diff changeset
   681
by (Blast_tac 1);
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   682
qed "Inter_iff";
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   683
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   684
Addsimps [Inter_iff];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   685
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   686
val prems = Goalw [Inter_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   687
    "[| !!X. X:C ==> A:X |] ==> A : Inter(C)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   688
by (REPEAT (ares_tac ([INT_I] @ prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   689
qed "InterI";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   690
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   691
(*A "destruct" rule -- every X in C contains A as an element, but
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   692
  A:X can hold when X:C does not!  This rule is analogous to "spec". *)
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   693
Goal "[| A : Inter(C);  X:C |] ==> A:X";
4477
b3e5857d8d99 New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents: 4469
diff changeset
   694
by Auto_tac;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   695
qed "InterD";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   696
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   697
(*"Classical" elimination rule -- does not require proving X:C *)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   698
val major::prems = Goalw [Inter_def]
2721
f08042e18c7d New version of InterE, like its ZF counterpart
paulson
parents: 2608
diff changeset
   699
    "[| A : Inter(C);  X~:C ==> R;  A:X ==> R |] ==> R";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   700
by (rtac (major RS INT_E) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   701
by (REPEAT (eresolve_tac prems 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   702
qed "InterE";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   703
2499
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   704
AddSIs [InterI];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   705
AddEs  [InterD, InterE];
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   706
0bc87b063447 Tidying of proofs. New theorems are enterred immediately into the
paulson
parents: 2031
diff changeset
   707
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   708
(*** Image of a set under a function ***)
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   709
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   710
(*Frequently b does not have the syntactic form of f(x).*)
10832
e33b47e4246d `` -> and ``` -> ``
nipkow
parents: 10482
diff changeset
   711
Goalw [image_def] "[| b=f(x);  x:A |] ==> b : f`A";
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   712
by (Blast_tac 1);
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   713
qed "image_eqI";
3909
e48e2fb8b895 Added image_eqI to simpset.
nipkow
parents: 3905
diff changeset
   714
Addsimps [image_eqI];
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   715
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   716
bind_thm ("imageI", refl RS image_eqI);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   717
8025
61dde9078e24 new theorem rev_image_eqI
paulson
parents: 8005
diff changeset
   718
(*This version's more effective when we already have the required x*)
10832
e33b47e4246d `` -> and ``` -> ``
nipkow
parents: 10482
diff changeset
   719
Goalw [image_def] "[| x:A;  b=f(x) |] ==> b : f`A";
8025
61dde9078e24 new theorem rev_image_eqI
paulson
parents: 8005
diff changeset
   720
by (Blast_tac 1);
61dde9078e24 new theorem rev_image_eqI
paulson
parents: 8005
diff changeset
   721
qed "rev_image_eqI";
61dde9078e24 new theorem rev_image_eqI
paulson
parents: 8005
diff changeset
   722
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   723
(*The eta-expansion gives variable-name preservation.*)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   724
val major::prems = Goalw [image_def]
10832
e33b47e4246d `` -> and ``` -> ``
nipkow
parents: 10482
diff changeset
   725
    "[| b : (%x. f(x))`A;  !!x.[| b=f(x);  x:A |] ==> P |] ==> P"; 
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   726
by (rtac (major RS CollectD RS bexE) 1);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   727
by (REPEAT (ares_tac prems 1));
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   728
qed "imageE";
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   729
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   730
AddIs  [image_eqI];
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   731
AddSEs [imageE]; 
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   732
10832
e33b47e4246d `` -> and ``` -> ``
nipkow
parents: 10482
diff changeset
   733
Goal "f`(A Un B) = f`A Un f`B";
2935
998cb95fdd43 Yet more fast_tac->blast_tac, and other tidying
paulson
parents: 2912
diff changeset
   734
by (Blast_tac 1);
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   735
qed "image_Un";
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   736
10832
e33b47e4246d `` -> and ``` -> ``
nipkow
parents: 10482
diff changeset
   737
Goal "(z : f`A) = (EX x:A. z = f x)";
3960
7a38fae985f9 New rewrite rules image_iff
paulson
parents: 3919
diff changeset
   738
by (Blast_tac 1);
7a38fae985f9 New rewrite rules image_iff
paulson
parents: 3919
diff changeset
   739
qed "image_iff";
7a38fae985f9 New rewrite rules image_iff
paulson
parents: 3919
diff changeset
   740
4523
16f5efe9812d New rule: image_subset
paulson
parents: 4510
diff changeset
   741
(*This rewrite rule would confuse users if made default.*)
10832
e33b47e4246d `` -> and ``` -> ``
nipkow
parents: 10482
diff changeset
   742
Goal "(f`A <= B) = (ALL x:A. f(x): B)";
4523
16f5efe9812d New rule: image_subset
paulson
parents: 4510
diff changeset
   743
by (Blast_tac 1);
16f5efe9812d New rule: image_subset
paulson
parents: 4510
diff changeset
   744
qed "image_subset_iff";
16f5efe9812d New rule: image_subset
paulson
parents: 4510
diff changeset
   745
11007
438f31613093 added diff_single_insert and subset_image_iff
oheimb
parents: 10832
diff changeset
   746
Goal "B <= f ` A = (? AA. AA <= A & B = f ` AA)";
438f31613093 added diff_single_insert and subset_image_iff
oheimb
parents: 10832
diff changeset
   747
by Safe_tac;
438f31613093 added diff_single_insert and subset_image_iff
oheimb
parents: 10832
diff changeset
   748
by  (Fast_tac 2);
438f31613093 added diff_single_insert and subset_image_iff
oheimb
parents: 10832
diff changeset
   749
by (res_inst_tac [("x","{a. a : A & f a : B}")] exI 1);
438f31613093 added diff_single_insert and subset_image_iff
oheimb
parents: 10832
diff changeset
   750
by (Fast_tac 1);
438f31613093 added diff_single_insert and subset_image_iff
oheimb
parents: 10832
diff changeset
   751
qed "subset_image_iff";
438f31613093 added diff_single_insert and subset_image_iff
oheimb
parents: 10832
diff changeset
   752
4523
16f5efe9812d New rule: image_subset
paulson
parents: 4510
diff changeset
   753
(*Replaces the three steps subsetI, imageE, hyp_subst_tac, but breaks too
16f5efe9812d New rule: image_subset
paulson
parents: 4510
diff changeset
   754
  many existing proofs.*)
10832
e33b47e4246d `` -> and ``` -> ``
nipkow
parents: 10482
diff changeset
   755
val prems = Goal "(!!x. x:A ==> f(x) : B) ==> f`A <= B";
4510
a37f515a0b25 New theorem image_subsetI
paulson
parents: 4477
diff changeset
   756
by (blast_tac (claset() addIs prems) 1);
a37f515a0b25 New theorem image_subsetI
paulson
parents: 4477
diff changeset
   757
qed "image_subsetI";
a37f515a0b25 New theorem image_subsetI
paulson
parents: 4477
diff changeset
   758
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   759
(*** Range of a function -- just a translation for image! ***)
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   760
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   761
Goal "b=f(x) ==> b : range(f)";
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   762
by (EVERY1 [etac image_eqI, rtac UNIV_I]);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   763
bind_thm ("range_eqI", UNIV_I RSN (2,image_eqI));
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   764
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   765
bind_thm ("rangeI", UNIV_I RS imageI);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   766
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5305
diff changeset
   767
val [major,minor] = Goal 
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3718
diff changeset
   768
    "[| b : range(%x. f(x));  !!x. b=f(x) ==> P |] ==> P"; 
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   769
by (rtac (major RS imageE) 1);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   770
by (etac minor 1);
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   771
qed "rangeE";
10482
41de88cb2108 UNIV_witness;
wenzelm
parents: 10233
diff changeset
   772
AddXEs [rangeE];
2912
3fac3e8d5d3e moved inj and surj from Set to Fun and Inv -> inv.
nipkow
parents: 2891
diff changeset
   773
1776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   774
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   775
(*** Set reasoning tools ***)
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   777
3912
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   778
(** Rewrite rules for boolean case-splitting: faster than 
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4770
diff changeset
   779
	addsplits[split_if]
3912
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   780
**)
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   781
4830
bd73675adbed Added a few lemmas.
nipkow
parents: 4770
diff changeset
   782
bind_thm ("split_if_eq1", read_instantiate [("P", "%x. x = ?b")] split_if);
bd73675adbed Added a few lemmas.
nipkow
parents: 4770
diff changeset
   783
bind_thm ("split_if_eq2", read_instantiate [("P", "%x. ?a = x")] split_if);
3912
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   784
5237
aebc63048f2d Deleted the redundant rule mem_if
paulson
parents: 5148
diff changeset
   785
(*Split ifs on either side of the membership relation.
aebc63048f2d Deleted the redundant rule mem_if
paulson
parents: 5148
diff changeset
   786
	Not for Addsimps -- can cause goals to blow up!*)
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9892
diff changeset
   787
bind_thm ("split_if_mem1", inst "P" "%x. x : ?b" split_if);
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9892
diff changeset
   788
bind_thm ("split_if_mem2", inst "P" "%x. ?a : x" split_if);
3912
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   789
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9088
diff changeset
   790
bind_thms ("split_ifs", [if_bool_eq_conj, split_if_eq1, split_if_eq2,
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9892
diff changeset
   791
			 split_if_mem1, split_if_mem2]);
3912
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   792
4ed64ad7fb42 New rewrite rules for simplifying conditionals
paulson
parents: 3909
diff changeset
   793
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4059
diff changeset
   794
(*Each of these has ALREADY been added to simpset() above.*)
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9088
diff changeset
   795
bind_thms ("mem_simps", [insert_iff, empty_iff, Un_iff, Int_iff, Compl_iff, Diff_iff, 
9fff97d29837 bind_thm(s);
wenzelm
parents: 9088
diff changeset
   796
                 mem_Collect_eq, UN_iff, Union_iff, INT_iff, Inter_iff]);
1776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   797
9041
paulson
parents: 8913
diff changeset
   798
(*Would like to add these, but the existing code only searches for the 
paulson
parents: 8913
diff changeset
   799
  outer-level constant, which in this case is just "op :"; we instead need
paulson
parents: 8913
diff changeset
   800
  to use term-nets to associate patterns with rules.  Also, if a rule fails to
paulson
parents: 8913
diff changeset
   801
  apply, then the formula should be kept.
paulson
parents: 8913
diff changeset
   802
  [("uminus", Compl_iff RS iffD1), ("op -", [Diff_iff RS iffD1]), 
paulson
parents: 8913
diff changeset
   803
   ("op Int", [IntD1,IntD2]),
paulson
parents: 8913
diff changeset
   804
   ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])]
paulson
parents: 8913
diff changeset
   805
 *)
paulson
parents: 8913
diff changeset
   806
val mksimps_pairs =
paulson
parents: 8913
diff changeset
   807
  [("Ball",[bspec])] @ mksimps_pairs;
1776
d7e77cb8ce5c moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents: 1762
diff changeset
   808
6291
2c3f72d9f5d1 simpler proofs of congruence rules
paulson
parents: 6171
diff changeset
   809
simpset_ref() := simpset() setmksimps (mksimps mksimps_pairs);
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   810
5256
e6983301ce5e Renamed equals0D to equals0E
paulson
parents: 5237
diff changeset
   811
Addsimps[subset_UNIV, subset_refl];
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   812
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   813
8001
14c8843cd35b new psubset lemma
paulson
parents: 7969
diff changeset
   814
(*** The 'proper subset' relation (<) ***)
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   815
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4830
diff changeset
   816
Goalw [psubset_def] "!!A::'a set. [| A <= B; A ~= B |] ==> A<B";
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   817
by (Blast_tac 1);
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   818
qed "psubsetI";
8913
0bc13d5e60b8 psubsetI is a safe rule
paulson
parents: 8839
diff changeset
   819
AddSIs [psubsetI];
3222
726a9b069947 Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents: 2935
diff changeset
   820
9088
453996655ac2 replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents: 9075
diff changeset
   821
Goalw [psubset_def]
453996655ac2 replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents: 9075
diff changeset
   822
  "(A < insert x B) = (if x:B then A<B else if x:A then A-{x} < B else A<=B)";
453996655ac2 replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents: 9075
diff changeset
   823
by (asm_simp_tac (simpset() addsimps [subset_insert_iff]) 1);
453996655ac2 replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents: 9075
diff changeset
   824
by (Blast_tac 1); 
453996655ac2 replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents: 9075
diff changeset
   825
qed "psubset_insert_iff";
4059
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   826
59c1422c9da5 New Blast_tac (and minor tidying...)
paulson
parents: 3960
diff changeset
   827
bind_thm ("psubset_eq", psubset_def RS meta_eq_to_obj_eq);
6443
6d5d3ecedf50 lemmas about proper subset relation;
wenzelm
parents: 6394
diff changeset
   828
6d5d3ecedf50 lemmas about proper subset relation;
wenzelm
parents: 6394
diff changeset
   829
bind_thm ("psubset_imp_subset", psubset_eq RS iffD1 RS conjunct1);
6d5d3ecedf50 lemmas about proper subset relation;
wenzelm
parents: 6394
diff changeset
   830
6d5d3ecedf50 lemmas about proper subset relation;
wenzelm
parents: 6394
diff changeset
   831
Goal"[| (A::'a set) < B; B <= C |] ==> A < C";
6d5d3ecedf50 lemmas about proper subset relation;
wenzelm
parents: 6394
diff changeset
   832
by (auto_tac (claset(), simpset() addsimps [psubset_eq]));
6d5d3ecedf50 lemmas about proper subset relation;
wenzelm
parents: 6394
diff changeset
   833
qed "psubset_subset_trans";
6d5d3ecedf50 lemmas about proper subset relation;
wenzelm
parents: 6394
diff changeset
   834
6d5d3ecedf50 lemmas about proper subset relation;
wenzelm
parents: 6394
diff changeset
   835
Goal"[| (A::'a set) <= B; B < C|] ==> A < C";
6d5d3ecedf50 lemmas about proper subset relation;
wenzelm
parents: 6394
diff changeset
   836
by (auto_tac (claset(), simpset() addsimps [psubset_eq]));
6d5d3ecedf50 lemmas about proper subset relation;
wenzelm
parents: 6394
diff changeset
   837
qed "subset_psubset_trans";
7717
e7ecfa617443 Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents: 7658
diff changeset
   838
8001
14c8843cd35b new psubset lemma
paulson
parents: 7969
diff changeset
   839
Goalw [psubset_def] "A < B ==> EX b. b : (B - A)";
14c8843cd35b new psubset lemma
paulson
parents: 7969
diff changeset
   840
by (Blast_tac 1);
14c8843cd35b new psubset lemma
paulson
parents: 7969
diff changeset
   841
qed "psubset_imp_ex_mem";
14c8843cd35b new psubset lemma
paulson
parents: 7969
diff changeset
   842
7717
e7ecfa617443 Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents: 7658
diff changeset
   843
9892
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   844
(* rulify setup *)
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   845
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   846
Goal "(!!x. x:A ==> P x) == Trueprop (ALL x:A. P x)";
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   847
by (simp_tac (simpset () addsimps (Ball_def :: thms "atomize")) 1);
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   848
qed "ball_eq";
7717
e7ecfa617443 Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents: 7658
diff changeset
   849
e7ecfa617443 Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents: 7658
diff changeset
   850
local
9892
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   851
  val ss = HOL_basic_ss addsimps
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   852
    (Drule.norm_hhf_eq :: map Thm.symmetric (ball_eq :: thms "atomize"));
7717
e7ecfa617443 Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents: 7658
diff changeset
   853
in
e7ecfa617443 Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents: 7658
diff changeset
   854
9892
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   855
structure Rulify = RulifyFun
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   856
 (val make_meta = Simplifier.simplify ss
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   857
  val full_make_meta = Simplifier.full_simplify ss);
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   858
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   859
structure BasicRulify: BASIC_RULIFY = Rulify;
be0389a64ce8 updated rulify setup;
wenzelm
parents: 9769
diff changeset
   860
open BasicRulify;
7717
e7ecfa617443 Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents: 7658
diff changeset
   861
e7ecfa617443 Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents: 7658
diff changeset
   862
end;