src/ZF/upair.ML
author paulson
Fri, 18 Feb 2000 15:35:29 +0100
changeset 8255 38f96394c099
parent 8183 344888de76c4
child 8551 5c22595bc599
permissions -rw-r--r--
new distributive laws
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1017
diff changeset
     1
(*  Title:      ZF/upair
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1017
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
UNORDERED pairs in Zermelo-Fraenkel Set Theory 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
Observe the order of dependence:
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
    Upair is defined in terms of Replace
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
    Un is defined in terms of Upair and Union (similarly for Int)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
    cons is defined in terms of Upair and Un
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
    Ordered pairs and descriptions are defined using cons ("set notation")
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
(*** Lemmas about power sets  ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1017
diff changeset
    17
val Pow_bottom = empty_subsetI RS PowI;         (* 0 : Pow(B) *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1017
diff changeset
    18
val Pow_top = subset_refl RS PowI;              (* A : Pow(A) *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
(*** Unordered pairs - Upair ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    23
qed_goalw "Upair_iff" thy [Upair_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
    "c : Upair(a,b) <-> (c=a | c=b)"
5468
paulson
parents: 5325
diff changeset
    25
 (fn _ => [ (Blast_tac 1) ]);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    26
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    27
Addsimps [Upair_iff];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    29
qed_goal "UpairI1" thy "a : Upair(a,b)"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    30
 (fn _ => [ Simp_tac 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    32
qed_goal "UpairI2" thy "b : Upair(a,b)"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    33
 (fn _ => [ Simp_tac 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    35
qed_goal "UpairE" thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
    "[| a : Upair(b,c);  a=b ==> P;  a=c ==> P |] ==> P"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
 (fn major::prems=>
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    38
  [ (rtac (major RS (Upair_iff RS iffD1 RS disjE)) 1),
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
    (REPEAT (eresolve_tac prems 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    41
AddSIs [UpairI1,UpairI2];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    42
AddSEs [UpairE];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    43
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    44
(*** Rules for binary union -- Un -- defined via Upair ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    46
qed_goalw "Un_iff" thy [Un_def] "c : A Un B <-> (c:A | c:B)"
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2690
diff changeset
    47
 (fn _ => [ Blast_tac 1 ]);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    48
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    49
Addsimps [Un_iff];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    51
qed_goal "UnI1" thy "!!c. c : A ==> c : A Un B"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    52
 (fn _ => [ Asm_simp_tac 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    54
qed_goal "UnI2" thy "!!c. c : B ==> c : A Un B"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    55
 (fn _ => [ Asm_simp_tac 1 ]);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    56
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    57
qed_goal "UnE" thy 
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
    "[| c : A Un B;  c:A ==> P;  c:B ==> P |] ==> P"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
 (fn major::prems=>
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    60
  [ (rtac (major RS (Un_iff RS iffD1 RS disjE)) 1),
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    61
    (REPEAT (eresolve_tac prems 1)) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
572
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
    63
(*Stronger version of the rule above*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    64
qed_goal "UnE'" thy
572
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
    65
    "[| c : A Un B;  c:A ==> P;  [| c:B;  c~:A |] ==> P |] ==> P"
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
    66
 (fn major::prems =>
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
    67
  [(rtac (major RS UnE) 1),
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
    68
   (eresolve_tac prems 1),
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
    69
   (rtac classical 1),
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
    70
   (eresolve_tac prems 1),
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
    71
   (swap_res_tac prems 1),
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
    72
   (etac notnotD 1)]);
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
    73
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    74
(*Classical introduction rule: no commitment to A vs B*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    75
qed_goal "UnCI" thy "(c ~: B ==> c : A) ==> c : A Un B"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    76
 (fn prems=>
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3914
diff changeset
    77
  [ Simp_tac 1, blast_tac (claset() addSIs prems) 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    79
AddSIs [UnCI];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    80
AddSEs [UnE];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    83
(*** Rules for small intersection -- Int -- defined via Upair ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    85
qed_goalw "Int_iff" thy [Int_def] "c : A Int B <-> (c:A & c:B)"
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2690
diff changeset
    86
 (fn _ => [ Blast_tac 1 ]);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    87
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    88
Addsimps [Int_iff];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    89
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    90
qed_goal "IntI" thy "!!c. [| c : A;  c : B |] ==> c : A Int B"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    91
 (fn _ => [ Asm_simp_tac 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    93
qed_goal "IntD1" thy "!!c. c : A Int B ==> c : A"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    94
 (fn _ => [ Asm_full_simp_tac 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    96
qed_goal "IntD2" thy "!!c. c : A Int B ==> c : B"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
    97
 (fn _ => [ Asm_full_simp_tac 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
    99
qed_goal "IntE" thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
    "[| c : A Int B;  [| c:A; c:B |] ==> P |] ==> P"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
 (fn prems=>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
  [ (resolve_tac prems 1),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
    (REPEAT (resolve_tac (prems RL [IntD1,IntD2]) 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   105
AddSIs [IntI];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   106
AddSEs [IntE];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   108
(*** Rules for set difference -- defined via Upair ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   110
qed_goalw "Diff_iff" thy [Diff_def] "c : A-B <-> (c:A & c~:B)"
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2690
diff changeset
   111
 (fn _ => [ Blast_tac 1 ]);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   112
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   113
Addsimps [Diff_iff];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   114
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   115
qed_goal "DiffI" thy "!!c. [| c : A;  c ~: B |] ==> c : A - B"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   116
 (fn _ => [ Asm_simp_tac 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   117
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   118
qed_goal "DiffD1" thy "!!c. c : A - B ==> c : A"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   119
 (fn _ => [ Asm_full_simp_tac 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   120
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   121
qed_goal "DiffD2" thy "!!c. c : A - B ==> c ~: B"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   122
 (fn _ => [ Asm_full_simp_tac 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   123
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   124
qed_goal "DiffE" thy
37
cebe01deba80 added ~: for "not in"
lcp
parents: 14
diff changeset
   125
    "[| c : A - B;  [| c:A; c~:B |] ==> P |] ==> P"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   126
 (fn prems=>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   127
  [ (resolve_tac prems 1),
485
5e00a676a211 Axiom of choice, cardinality results, etc.
lcp
parents: 437
diff changeset
   128
    (REPEAT (ares_tac (prems RL [DiffD1, DiffD2]) 1)) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   129
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   130
AddSIs [DiffI];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   131
AddSEs [DiffE];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   133
(*** Rules for cons -- defined via Un and Upair ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   135
qed_goalw "cons_iff" thy [cons_def] "a : cons(b,A) <-> (a=b | a:A)"
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2690
diff changeset
   136
 (fn _ => [ Blast_tac 1 ]);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   137
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   138
Addsimps [cons_iff];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   139
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   140
qed_goal "consI1" thy "a : cons(a,B)"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   141
 (fn _ => [ Simp_tac 1 ]);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   142
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   143
Addsimps [consI1];
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6068
diff changeset
   144
AddTCs   [consI1];   (*risky as a typechecking rule, but solves otherwise
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6068
diff changeset
   145
		       unconstrained goals of the form  x : ?A*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   146
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   147
qed_goal "consI2" thy "!!B. a : B ==> a : cons(b,B)"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   148
 (fn _ => [ Asm_simp_tac 1 ]);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   149
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   150
qed_goal "consE" thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   151
    "[| a : cons(b,A);  a=b ==> P;  a:A ==> P |] ==> P"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   152
 (fn major::prems=>
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   153
  [ (rtac (major RS (cons_iff RS iffD1 RS disjE)) 1),
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   154
    (REPEAT (eresolve_tac (prems @ [UpairE]) 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   155
572
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
   156
(*Stronger version of the rule above*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   157
qed_goal "consE'" thy
572
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
   158
    "[| a : cons(b,A);  a=b ==> P;  [| a:A;  a~=b |] ==> P |] ==> P"
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
   159
 (fn major::prems =>
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
   160
  [(rtac (major RS consE) 1),
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
   161
   (eresolve_tac prems 1),
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
   162
   (rtac classical 1),
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
   163
   (eresolve_tac prems 1),
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
   164
   (swap_res_tac prems 1),
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
   165
   (etac notnotD 1)]);
13c30ac40f8f ZF/upair/consE', UnE': new
lcp
parents: 485
diff changeset
   166
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   167
(*Classical introduction rule*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   168
qed_goal "consCI" thy "(a~:B ==> a=b) ==> a: cons(b,B)"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   169
 (fn prems=>
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3914
diff changeset
   170
  [ Simp_tac 1, blast_tac (claset() addSIs prems) 1 ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   171
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   172
AddSIs [consCI];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   173
AddSEs [consE];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   174
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   175
qed_goal "cons_not_0" thy "cons(a,B) ~= 0"
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3914
diff changeset
   176
 (fn _ => [ (blast_tac (claset() addEs [equalityE]) 1) ]);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   177
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   178
bind_thm ("cons_neq_0", cons_not_0 RS notE);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   179
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   180
Addsimps [cons_not_0, cons_not_0 RS not_sym];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   181
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   182
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   183
(*** Singletons - using cons ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   184
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   185
qed_goal "singleton_iff" thy "a : {b} <-> a=b"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   186
 (fn _ => [ Simp_tac 1 ]);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   187
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   188
qed_goal "singletonI" thy "a : {a}"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   189
 (fn _=> [ (rtac consI1 1) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   190
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   191
bind_thm ("singletonE", make_elim (singleton_iff RS iffD1));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   192
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   193
AddSIs [singletonI];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   194
AddSEs [singletonE];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   195
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   196
(*** Rules for Descriptions ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   197
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   198
qed_goalw "the_equality" thy [the_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   199
    "[| P(a);  !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a"
738
3054a10ed5b5 the_equality: more careful use of addSIs and addIs
lcp
parents: 686
diff changeset
   200
 (fn [pa,eq] =>
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3914
diff changeset
   201
  [ (fast_tac (claset() addSIs [pa] addEs [eq RS subst]) 1) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   202
5506
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   203
AddIs [the_equality];
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   204
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   205
(* Only use this if you already know EX!x. P(x) *)
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6163
diff changeset
   206
Goal "[| EX! x. P(x);  P(a) |] ==> (THE x. P(x)) = a";
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6163
diff changeset
   207
by (Blast_tac 1);
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6163
diff changeset
   208
qed "the_equality2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   209
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6163
diff changeset
   210
Goal "EX! x. P(x) ==> P(THE x. P(x))";
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6163
diff changeset
   211
by (etac ex1E 1);
8183
344888de76c4 expandshort
paulson
parents: 8127
diff changeset
   212
by (stac the_equality 1);
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6163
diff changeset
   213
by (REPEAT (Blast_tac 1));
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6163
diff changeset
   214
qed "theI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   215
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 37
diff changeset
   216
(*the_cong is no longer necessary: if (ALL y.P(y)<->Q(y)) then 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 37
diff changeset
   217
  (THE x.P(x))  rewrites to  (THE x. Q(x))  *)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 37
diff changeset
   218
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 37
diff changeset
   219
(*If it's "undefined", it's zero!*)
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6163
diff changeset
   220
Goalw  [the_def] "!!P. ~ (EX! x. P(x)) ==> (THE x. P(x))=0";
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6163
diff changeset
   221
by (blast_tac (claset() addSEs [ReplaceE]) 1);
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6163
diff changeset
   222
qed "the_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   223
5506
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   224
(*Easier to apply than theI: conclusion has only one occurrence of P*)
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   225
val prems = 
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   226
Goal "[| ~ Q(0) ==> EX! x. P(x);  !!x. P(x) ==> Q(x) |] ==> Q(THE x. P(x))";
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   227
by (rtac classical 1);
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   228
by (resolve_tac prems 1);
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   229
by (rtac theI 1);
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   230
by (rtac classical 1);
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   231
by (resolve_tac prems 1);
6163
be8234f37e48 expandshort
paulson
parents: 6153
diff changeset
   232
by (etac (the_0 RS subst) 1);
be8234f37e48 expandshort
paulson
parents: 6153
diff changeset
   233
by (assume_tac 1);
5506
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   234
qed "theI2";
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   235
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   236
(*** if -- a conditional expression for formulae ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   237
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5506
diff changeset
   238
Goalw [if_def] "(if True then a else b) = a";
5506
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   239
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 738
diff changeset
   240
qed "if_true";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   241
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5506
diff changeset
   242
Goalw [if_def] "(if False then a else b) = b";
5506
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   243
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 738
diff changeset
   244
qed "if_false";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   245
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   246
(*Never use with case splitting, or if P is known to be true or false*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   247
val prems = Goalw [if_def]
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5506
diff changeset
   248
    "[| P<->Q;  Q ==> a=c;  ~Q ==> b=d |] \
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5506
diff changeset
   249
\    ==> (if P then a else b) = (if Q then c else d)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3914
diff changeset
   250
by (simp_tac (simpset() addsimps prems addcongs [conj_cong]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 738
diff changeset
   251
qed "if_cong";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   252
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   253
(*Not needed for rewriting, since P would rewrite to True anyway*)
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5506
diff changeset
   254
Goalw [if_def] "P ==> (if P then a else b) = a";
5506
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   255
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 738
diff changeset
   256
qed "if_P";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   257
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   258
(*Not needed for rewriting, since P would rewrite to False anyway*)
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5506
diff changeset
   259
Goalw [if_def] "~P ==> (if P then a else b) = b";
5506
e07254044384 stronger version of theI2
paulson
parents: 5488
diff changeset
   260
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 738
diff changeset
   261
qed "if_not_P";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   262
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   263
Addsimps [if_true, if_false];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   264
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   265
qed_goal "split_if" thy
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5506
diff changeset
   266
    "P(if Q then x else y) <-> ((Q --> P(x)) & (~Q --> P(y)))"
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   267
 (fn _=> [ (case_tac "Q" 1),
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   268
           (Asm_simp_tac 1),
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   269
           (Asm_simp_tac 1) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   270
3914
9e393b363c71 New rewrite rules for simplifying conditionals
paulson
parents: 3840
diff changeset
   271
(** Rewrite rules for boolean case-splitting: faster than 
5116
8eb343ab5748 Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents: 4091
diff changeset
   272
	addsplits[split_if]
3914
9e393b363c71 New rewrite rules for simplifying conditionals
paulson
parents: 3840
diff changeset
   273
**)
9e393b363c71 New rewrite rules for simplifying conditionals
paulson
parents: 3840
diff changeset
   274
5116
8eb343ab5748 Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents: 4091
diff changeset
   275
bind_thm ("split_if_eq1", read_instantiate [("P", "%x. x = ?b")] split_if);
8eb343ab5748 Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents: 4091
diff changeset
   276
bind_thm ("split_if_eq2", read_instantiate [("P", "%x. ?a = x")] split_if);
3914
9e393b363c71 New rewrite rules for simplifying conditionals
paulson
parents: 3840
diff changeset
   277
5116
8eb343ab5748 Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents: 4091
diff changeset
   278
bind_thm ("split_if_mem1", read_instantiate [("P", "%x. x : ?b")] split_if);
8eb343ab5748 Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents: 4091
diff changeset
   279
bind_thm ("split_if_mem2", read_instantiate [("P", "%x. ?a : x")] split_if);
3914
9e393b363c71 New rewrite rules for simplifying conditionals
paulson
parents: 3840
diff changeset
   280
5116
8eb343ab5748 Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents: 4091
diff changeset
   281
val split_ifs = [split_if_eq1, split_if_eq2,
8eb343ab5748 Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents: 4091
diff changeset
   282
		 split_if_mem1, split_if_mem2];
3914
9e393b363c71 New rewrite rules for simplifying conditionals
paulson
parents: 3840
diff changeset
   283
5116
8eb343ab5748 Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents: 4091
diff changeset
   284
(*Logically equivalent to split_if_mem2*)
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5506
diff changeset
   285
qed_goal "if_iff" thy "a: (if P then x else y) <-> P & a:x | ~P & a:y"
5116
8eb343ab5748 Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents: 4091
diff changeset
   286
 (fn _=> [ (simp_tac (simpset() addsplits [split_if]) 1) ]);
1017
6a402dc505cf Proved if_iff and used it to simplify proof of if_type.
lcp
parents: 985
diff changeset
   287
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   288
qed_goal "if_type" thy
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5506
diff changeset
   289
    "[| P ==> a: A;  ~P ==> b: A |] ==> (if P then a else b): A"
1017
6a402dc505cf Proved if_iff and used it to simplify proof of if_type.
lcp
parents: 985
diff changeset
   290
 (fn prems=> [ (simp_tac 
5116
8eb343ab5748 Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents: 4091
diff changeset
   291
                (simpset() addsimps prems addsplits [split_if]) 1) ]);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6068
diff changeset
   292
AddTCs [if_type];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   293
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   294
(*** Foundation lemmas ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   295
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   296
(*was called mem_anti_sym*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   297
qed_goal "mem_asym" thy "[| a:b;  ~P ==> b:a |] ==> P"
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2690
diff changeset
   298
 (fn prems=>
6476784dba1c Converted to call blast_tac.
paulson
parents: 2690
diff changeset
   299
  [ (rtac classical 1),
6476784dba1c Converted to call blast_tac.
paulson
parents: 2690
diff changeset
   300
    (res_inst_tac [("A1","{a,b}")] (foundation RS disjE) 1),
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3914
diff changeset
   301
    REPEAT (blast_tac (claset() addIs prems addSEs [equalityE]) 1) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   302
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   303
(*was called mem_anti_refl*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   304
qed_goal "mem_irrefl" thy "a:a ==> P"
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   305
 (fn [major]=> [ (rtac ([major,major] MRS mem_asym) 1) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   306
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   307
(*mem_irrefl should NOT be added to default databases:
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   308
      it would be tried on most goals, making proofs slower!*)
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   309
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   310
qed_goal "mem_not_refl" thy "a ~: a"
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   311
 (K [ (rtac notI 1), (etac mem_irrefl 1) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   312
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 37
diff changeset
   313
(*Good for proving inequalities by rewriting*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   314
qed_goal "mem_imp_not_eq" thy "!!a A. a:A ==> a ~= A"
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3914
diff changeset
   315
 (fn _=> [ blast_tac (claset() addSEs [mem_irrefl]) 1 ]);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 37
diff changeset
   316
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   317
(*** Rules for succ ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   318
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   319
qed_goalw "succ_iff" thy [succ_def] "i : succ(j) <-> i=j | i:j"
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2690
diff changeset
   320
 (fn _ => [ Blast_tac 1 ]);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   321
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   322
qed_goalw "succI1" thy [succ_def] "i : succ(i)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   323
 (fn _=> [ (rtac consI1 1) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   324
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   325
Addsimps [succI1];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   326
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   327
qed_goalw "succI2" thy [succ_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   328
    "i : j ==> i : succ(j)"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   329
 (fn [prem]=> [ (rtac (prem RS consI2) 1) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   330
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   331
qed_goalw "succE" thy [succ_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   332
    "[| i : succ(j);  i=j ==> P;  i:j ==> P |] ==> P"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   333
 (fn major::prems=>
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   334
  [ (rtac (major RS consE) 1),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   335
    (REPEAT (eresolve_tac prems 1)) ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   336
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   337
(*Classical introduction rule*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   338
qed_goal "succCI" thy "(i~:j ==> i=j) ==> i: succ(j)"
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   339
 (fn [prem]=>
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   340
  [ (rtac (disjCI RS (succ_iff RS iffD2)) 1),
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   341
    (etac prem 1) ]);
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   342
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   343
AddSIs [succCI];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   344
AddSEs [succE];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   345
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   346
qed_goal "succ_not_0" thy "succ(n) ~= 0"
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3914
diff changeset
   347
 (fn _=> [ (blast_tac (claset() addSEs [equalityE]) 1) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   348
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   349
bind_thm ("succ_neq_0", succ_not_0 RS notE);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   350
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   351
Addsimps [succ_not_0, succ_not_0 RS not_sym];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   352
AddSEs [succ_neq_0, sym RS succ_neq_0];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   353
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   354
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   355
(* succ(c) <= B ==> c : B *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   356
val succ_subsetD = succI1 RSN (2,subsetD);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   357
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   358
(* succ(b) ~= b *)
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   359
bind_thm ("succ_neq_self", succI1 RS mem_imp_not_eq RS not_sym);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   360
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   361
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5242
diff changeset
   362
qed_goal "succ_inject_iff" thy "succ(m) = succ(n) <-> m=n"
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3914
diff changeset
   363
 (fn _=> [ (blast_tac (claset() addEs [mem_asym] addSEs [equalityE]) 1) ]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   364
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   365
bind_thm ("succ_inject", succ_inject_iff RS iffD1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   366
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   367
Addsimps [succ_inject_iff];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   368
AddSDs [succ_inject];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   369
2877
6476784dba1c Converted to call blast_tac.
paulson
parents: 2690
diff changeset
   370
(*Not needed now that cons is available.  Deletion reduces the search space.*)
6476784dba1c Converted to call blast_tac.
paulson
parents: 2690
diff changeset
   371
Delrules [UpairI1,UpairI2,UpairE];
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   372
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1609
diff changeset
   373
use"simpdata.ML";