src/ZF/ex/PropLog.ML
author wenzelm
Thu, 30 Nov 2000 20:05:34 +0100
changeset 10551 ec9fab41b3a0
parent 6154 6a00a5baef2b
child 11233 34c81a796ee3
permissions -rw-r--r--
renamed "equivalence_class" to "class";
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     1
(*  Title:      ZF/ex/prop-log.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     3
    Author:     Tobias Nipkow & Lawrence C Paulson
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
     6
Inductive definition of propositional logic.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
Soundness and completeness w.r.t. truth-tables.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
Prove: If H|=p then G|=p where G:Fin(H)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5618
diff changeset
    12
Addsimps prop.intrs;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
(*** Semantics of propositional logic ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
(** The function is_true **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    18
Goalw [is_true_def] "is_true(Fls,t) <-> False";
5618
paulson
parents: 5325
diff changeset
    19
by (Simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    20
qed "is_true_Fls";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    22
Goalw [is_true_def] "is_true(#v,t) <-> v:t";
5618
paulson
parents: 5325
diff changeset
    23
by (Simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    24
qed "is_true_Var";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
5618
paulson
parents: 5325
diff changeset
    26
Goalw [is_true_def] "is_true(p=>q,t) <-> (is_true(p,t)-->is_true(q,t))";
paulson
parents: 5325
diff changeset
    27
by (Simp_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    28
qed "is_true_Imp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5618
diff changeset
    30
Addsimps [is_true_Fls, is_true_Var, is_true_Imp];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
(*** Proof theory of propositional logic ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    35
Goalw thms.defs "G<=H ==> thms(G) <= thms(H)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
by (rtac lfp_mono 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    37
by (REPEAT (rtac thms.bnd_mono 1));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    39
qed "thms_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    41
val thms_in_pl = thms.dom_subset RS subsetD;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6068
diff changeset
    43
val ImpE = prop.mk_cases "p=>q : prop";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    44
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    45
(*Stronger Modus Ponens rule: no typechecking!*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    46
Goal "[| H |- p=>q;  H |- p |] ==> H |- q";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    47
by (rtac thms.MP 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
by (REPEAT (eresolve_tac [asm_rl, thms_in_pl, thms_in_pl RS ImpE] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    49
qed "thms_MP";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
(*Rule is called I for Identity Combinator, not for Introduction*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    52
Goal "p:prop ==> H |- p=>p";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    53
by (rtac (thms.S RS thms_MP RS thms_MP) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    54
by (rtac thms.K 5);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    55
by (rtac thms.K 4);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    56
by (REPEAT (ares_tac prop.intrs 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    57
qed "thms_I";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
(** Weakening, left and right **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    60
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
(* [| G<=H;  G|-p |] ==> H|-p   Order of premises is convenient with RS*)
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    62
bind_thm ("weaken_left", (thms_mono RS subsetD));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
(* H |- p ==> cons(a,H) |- p *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
val weaken_left_cons = subset_consI RS weaken_left;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
val weaken_left_Un1  = Un_upper1 RS weaken_left;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
val weaken_left_Un2  = Un_upper2 RS weaken_left;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    69
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    70
Goal "[| H |- q;  p:prop |] ==> H |- p=>q";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    71
by (rtac (thms.K RS thms_MP) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
by (REPEAT (ares_tac [thms_in_pl] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    73
qed "weaken_right";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
(*The deduction theorem*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    76
Goal "[| cons(p,H) |- q;  p:prop |] ==>  H |- p=>q";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    77
by (etac thms.induct 1);
6154
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
    78
by (blast_tac (claset() addIs [thms_I, thms.H RS weaken_right]) 1);
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
    79
by (blast_tac (claset() addIs [thms.K RS weaken_right]) 1);
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
    80
by (blast_tac (claset() addIs [thms.S RS weaken_right]) 1);
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
    81
by (blast_tac (claset() addIs [thms.DN RS weaken_right]) 1);
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
    82
by (blast_tac (claset() addIs [thms.S RS thms_MP RS thms_MP]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    83
qed "deduction";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    85
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
(*The cut rule*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    87
Goal "[| H|-p;  cons(p,H) |- q |] ==>  H |- q";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
by (rtac (deduction RS thms_MP) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
by (REPEAT (ares_tac [thms_in_pl] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    90
qed "cut";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    92
Goal "[| H |- Fls; p:prop |] ==> H |- p";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    93
by (rtac (thms.DN RS thms_MP) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
by (rtac weaken_right 2);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    95
by (REPEAT (ares_tac (prop.intrs@[consI1]) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    96
qed "thms_FlsE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
(* [| H |- p=>Fls;  H |- p;  q: prop |] ==> H |- q *)
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    99
bind_thm ("thms_notE", (thms_MP RS thms_FlsE));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
(*Soundness of the rules wrt truth-table semantics*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   102
Goalw [logcon_def] "H |- p ==> H |= p";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   103
by (etac thms.induct 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5618
diff changeset
   104
by Auto_tac;
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   105
qed "soundness";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
(*** Towards the completeness proof ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   108
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   109
val [premf,premq] = goal PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
    "[| H |- p=>Fls; q: prop |] ==> H |- p=>q";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   111
by (rtac (premf RS thms_in_pl RS ImpE) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   112
by (rtac deduction 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   113
by (rtac (premf RS weaken_left_cons RS thms_notE) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   114
by (REPEAT (ares_tac [premq, consI1, thms.H] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   115
qed "Fls_Imp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   116
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   117
val [premp,premq] = goal PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   118
    "[| H |- p;  H |- q=>Fls |] ==> H |- (p=>q)=>Fls";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
by (cut_facts_tac ([premp,premq] RL [thms_in_pl]) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   120
by (etac ImpE 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   121
by (rtac deduction 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   122
by (rtac (premq RS weaken_left_cons RS thms_MP) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   123
by (rtac (consI1 RS thms.H RS thms_MP) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
by (rtac (premp RS weaken_left_cons) 2);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   125
by (REPEAT (ares_tac prop.intrs 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   126
qed "Imp_Fls";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   127
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   128
(*Typical example of strengthening the induction formula*)
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 6046
diff changeset
   129
Goal "p: prop ==> hyps(p,t) |- (if is_true(p,t) then p else p=>Fls)";
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5618
diff changeset
   130
by (Simp_tac 1);
6154
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
   131
by (induct_tac "p" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2873
diff changeset
   132
by (ALLGOALS (asm_simp_tac (simpset() addsimps [thms_I, thms.H])));
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2873
diff changeset
   133
by (safe_tac (claset() addSEs [Fls_Imp RS weaken_left_Un1, 
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   134
			       Fls_Imp RS weaken_left_Un2]));
6154
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
   135
by (ALLGOALS (blast_tac (claset() addIs [weaken_left_Un1, weaken_left_Un2, 
5618
paulson
parents: 5325
diff changeset
   136
					weaken_right, Imp_Fls])));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   137
qed "hyps_thms_if";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   138
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   139
(*Key lemma for completeness; yields a set of assumptions satisfying p*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   140
Goalw [logcon_def] "[| p: prop;  0 |= p |] ==> hyps(p,t) |- p";
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   141
by (dtac hyps_thms_if 1);
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   142
by (Asm_full_simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   143
qed "logcon_thms_p";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   145
(*For proving certain theorems in our new propositional logic*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   146
val thms_cs = 
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   147
    ZF_cs addSIs (prop.intrs @ [deduction])
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   148
          addIs [thms_in_pl, thms.H, thms.H RS thms_MP];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   149
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   150
(*The excluded middle in the form of an elimination rule*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   151
Goal "[| p: prop;  q: prop |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   152
by (rtac (deduction RS deduction) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   153
by (rtac (thms.DN RS thms_MP) 1);
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   154
by (ALLGOALS (blast_tac thms_cs));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   155
qed "thms_excluded_middle";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   157
(*Hard to prove directly because it requires cuts*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   158
Goal "[| cons(p,H) |- q;  cons(p=>Fls,H) |- q;  p: prop |] ==> H |- q";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   159
by (rtac (thms_excluded_middle RS thms_MP RS thms_MP) 1);
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   160
by (REPEAT (ares_tac (prop.intrs@[deduction,thms_in_pl]) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   161
qed "thms_excluded_middle_rule";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   162
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   163
(*** Completeness -- lemmas for reducing the set of assumptions ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   165
(*For the case hyps(p,t)-cons(#v,Y) |- p;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
  we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   167
Goal "p: prop ==> hyps(p, t-{v}) <= cons(#v=>Fls, hyps(p,t)-{#v})";
6154
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
   168
by (induct_tac "p" 1);
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6068
diff changeset
   169
by Auto_tac;
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   170
qed "hyps_Diff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   171
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   172
(*For the case hyps(p,t)-cons(#v => Fls,Y) |- p;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
  we also have hyps(p,t)-{#v=>Fls} <= hyps(p, cons(v,t)) *)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   174
Goal "p: prop ==> hyps(p, cons(v,t)) <= cons(#v, hyps(p,t)-{#v=>Fls})";
6154
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
   175
by (induct_tac "p" 1);
6141
a6922171b396 removal of the (thm list) argument of mk_cases
paulson
parents: 6068
diff changeset
   176
by Auto_tac;
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   177
qed "hyps_cons";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   178
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   179
(** Two lemmas for use with weaken_left **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   180
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   181
Goal "B-C <= cons(a, B-cons(a,C))";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   182
by (Fast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   183
qed "cons_Diff_same";
0
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: 5137
diff changeset
   185
Goal "cons(a, B-{c}) - D <= cons(a, B-cons(c,D))";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   186
by (Fast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   187
qed "cons_Diff_subset2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   188
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   189
(*The set hyps(p,t) is finite, and elements have the form #v or #v=>Fls;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   190
 could probably prove the stronger hyps(p,t) : Fin(hyps(p,0) Un hyps(p,nat))*)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5137
diff changeset
   191
Goal "p: prop ==> hyps(p,t) : Fin(UN v:nat. {#v, #v=>Fls})";
6154
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
   192
by (induct_tac "p" 1);
5618
paulson
parents: 5325
diff changeset
   193
by (asm_simp_tac (simpset() addsimps [UN_I]) 2);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   194
by (ALLGOALS Asm_simp_tac);
6154
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
   195
by (blast_tac (claset() addIs Fin.intrs) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   196
qed "hyps_finite";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   197
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   198
val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   199
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   200
(*Induction on the finite set of assumptions hyps(p,t0).
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   201
  We may repeatedly subtract assumptions until none are left!*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   202
val [premp,sat] = goal PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   203
    "[| p: prop;  0 |= p |] ==> ALL t. hyps(p,t) - hyps(p,t0) |- p";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   204
by (rtac (premp RS hyps_finite RS Fin_induct) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2873
diff changeset
   205
by (simp_tac (simpset() addsimps [premp, sat, logcon_thms_p, Diff_0]) 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   206
by Safe_tac;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   207
(*Case hyps(p,t)-cons(#v,Y) |- p *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   208
by (rtac thms_excluded_middle_rule 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   209
by (etac prop.Var_I 3);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   210
by (rtac (cons_Diff_same RS weaken_left) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   211
by (etac spec 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   212
by (rtac (cons_Diff_subset2 RS weaken_left) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   213
by (rtac (premp RS hyps_Diff RS Diff_weaken_left) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   214
by (etac spec 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   215
(*Case hyps(p,t)-cons(#v => Fls,Y) |- p *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
by (rtac thms_excluded_middle_rule 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   217
by (etac prop.Var_I 3);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   218
by (rtac (cons_Diff_same RS weaken_left) 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   219
by (etac spec 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   220
by (rtac (cons_Diff_subset2 RS weaken_left) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   221
by (rtac (premp RS hyps_cons RS Diff_weaken_left) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   222
by (etac spec 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   223
qed "completeness_0_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   224
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   225
(*The base case for completeness*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   226
val [premp,sat] = goal PropLog.thy "[| p: prop;  0 |= p |] ==> 0 |- p";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   227
by (rtac (Diff_cancel RS subst) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   228
by (rtac (sat RS (premp RS completeness_0_lemma RS spec)) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   229
qed "completeness_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   230
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   231
(*A semantic analogue of the Deduction Theorem*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   232
Goalw [logcon_def] "[| cons(p,H) |= q |] ==> H |= p=>q";
6154
6a00a5baef2b automatic insertion of datatype intr rules into claset
paulson
parents: 6141
diff changeset
   233
by Auto_tac;
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   234
qed "logcon_Imp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   235
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   236
Goal "H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   237
by (etac Fin_induct 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2873
diff changeset
   238
by (safe_tac (claset() addSIs [completeness_0]));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   239
by (rtac (weaken_left_cons RS thms_MP) 1);
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   240
by (blast_tac (claset() addSIs (logcon_Imp::prop.intrs)) 1);
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   241
by (blast_tac thms_cs 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   242
qed "completeness_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   243
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   244
val completeness = completeness_lemma RS bspec RS mp;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   245
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   246
val [finite] = goal PropLog.thy "H: Fin(prop) ==> H |- p <-> H |= p & p:prop";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2873
diff changeset
   247
by (fast_tac (claset() addSEs [soundness, finite RS completeness, 
5618
paulson
parents: 5325
diff changeset
   248
			       thms_in_pl]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   249
qed "thms_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   250
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   251
writeln"Reached end of file.";