src/ZF/ex/PropLog.ML
author clasohm
Tue, 30 Jan 1996 13:42:57 +0100
changeset 1461 6bcb44e4d6e5
parent 782 200a16083201
child 2469 b50b8c0eec01
permissions -rw-r--r--
expanded tabs
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
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
For ex/prop-log.thy.  Inductive definition of propositional logic.
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
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
open PropLog;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
(*** prop_rec -- by Vset recursion ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
(** conversion rules **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
goal PropLog.thy "prop_rec(Fls,b,c,d) = b";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
by (rtac (prop_rec_def RS def_Vrec RS trans) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    20
by (rewrite_goals_tac prop.con_defs);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    21
by (simp_tac rank_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    22
qed "prop_rec_Fls";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
goal PropLog.thy "prop_rec(#v,b,c,d) = c(v)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
by (rtac (prop_rec_def RS def_Vrec RS trans) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    26
by (rewrite_goals_tac prop.con_defs);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    27
by (simp_tac rank_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    28
qed "prop_rec_Var";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
goal PropLog.thy "prop_rec(p=>q,b,c,d) = \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
\      d(p, q, prop_rec(p,b,c,d), prop_rec(q,b,c,d))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
by (rtac (prop_rec_def RS def_Vrec RS trans) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    33
by (rewrite_goals_tac prop.con_defs);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    34
by (simp_tac rank_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    35
qed "prop_rec_Imp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
val prop_rec_ss = 
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    38
    arith_ss addsimps [prop_rec_Fls, prop_rec_Var, prop_rec_Imp];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
(*** Semantics of propositional logic ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
(** The function is_true **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    44
goalw PropLog.thy [is_true_def] "is_true(Fls,t) <-> False";
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    45
by (simp_tac (prop_rec_ss addsimps [one_not_0 RS not_sym]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    46
qed "is_true_Fls";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
goalw PropLog.thy [is_true_def] "is_true(#v,t) <-> v:t";
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    49
by (simp_tac (prop_rec_ss addsimps [one_not_0 RS not_sym] 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    50
              setloop (split_tac [expand_if])) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    51
qed "is_true_Var";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
goalw PropLog.thy [is_true_def]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
    "is_true(p=>q,t) <-> (is_true(p,t)-->is_true(q,t))";
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    55
by (simp_tac (prop_rec_ss setloop (split_tac [expand_if])) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    56
qed "is_true_Imp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
(** The function hyps **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    60
goalw PropLog.thy [hyps_def] "hyps(Fls,t) = 0";
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    61
by (simp_tac prop_rec_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    62
qed "hyps_Fls";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
goalw PropLog.thy [hyps_def] "hyps(#v,t) = {if(v:t, #v, #v=>Fls)}";
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    65
by (simp_tac prop_rec_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    66
qed "hyps_Var";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
goalw PropLog.thy [hyps_def] "hyps(p=>q,t) = hyps(p,t) Un hyps(q,t)";
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    69
by (simp_tac prop_rec_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    70
qed "hyps_Imp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
val prop_ss = prop_rec_ss 
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    73
    addsimps prop.intrs
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    74
    addsimps [is_true_Fls, is_true_Var, is_true_Imp,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    75
              hyps_Fls, hyps_Var, hyps_Imp];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
(*** Proof theory of propositional logic ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    79
goalw PropLog.thy thms.defs "!!G H. G<=H ==> thms(G) <= thms(H)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
by (rtac lfp_mono 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    81
by (REPEAT (rtac thms.bnd_mono 1));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    83
qed "thms_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    85
val thms_in_pl = thms.dom_subset RS subsetD;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    87
val ImpE = prop.mk_cases prop.con_defs "p=>q : prop";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    89
(*Stronger Modus Ponens rule: no typechecking!*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    90
goal PropLog.thy "!!p q H. [| H |- p=>q;  H |- p |] ==> H |- q";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    91
by (rtac thms.MP 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
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
    93
qed "thms_MP";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
(*Rule is called I for Identity Combinator, not for Introduction*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    96
goal PropLog.thy "!!p H. p:prop ==> H |- p=>p";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    97
by (rtac (thms.S RS thms_MP RS thms_MP) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    98
by (rtac thms.K 5);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    99
by (rtac thms.K 4);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   100
by (REPEAT (ares_tac prop.intrs 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   101
qed "thms_I";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
(** Weakening, left and right **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
(* [| 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
   106
bind_thm ("weaken_left", (thms_mono RS subsetD));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   108
(* H |- p ==> cons(a,H) |- p *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
val weaken_left_cons = subset_consI RS weaken_left;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   111
val weaken_left_Un1  = Un_upper1 RS weaken_left;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   112
val weaken_left_Un2  = Un_upper2 RS weaken_left;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   113
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   114
goal PropLog.thy "!!H p q. [| H |- q;  p:prop |] ==> H |- p=>q";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   115
by (rtac (thms.K RS thms_MP) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   116
by (REPEAT (ares_tac [thms_in_pl] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   117
qed "weaken_right";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   118
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
(*The deduction theorem*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   120
goal PropLog.thy "!!p q H. [| cons(p,H) |- q;  p:prop |] ==>  H |- p=>q";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   121
by (etac thms.induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   122
by (fast_tac (ZF_cs addIs [thms_I, thms.H RS weaken_right]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   123
by (fast_tac (ZF_cs addIs [thms.K RS weaken_right]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   124
by (fast_tac (ZF_cs addIs [thms.S RS weaken_right]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   125
by (fast_tac (ZF_cs addIs [thms.DN RS weaken_right]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   126
by (fast_tac (ZF_cs addIs [thms.S RS thms_MP RS thms_MP]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   127
qed "deduction";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   128
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   129
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
(*The cut rule*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   131
goal PropLog.thy "!!H p q. [| H|-p;  cons(p,H) |- q |] ==>  H |- q";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
by (rtac (deduction RS thms_MP) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   133
by (REPEAT (ares_tac [thms_in_pl] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   134
qed "cut";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   135
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   136
goal PropLog.thy "!!H p. [| H |- Fls; p:prop |] ==> H |- p";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   137
by (rtac (thms.DN RS thms_MP) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   138
by (rtac weaken_right 2);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   139
by (REPEAT (ares_tac (prop.intrs@[consI1]) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   140
qed "thms_FlsE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   141
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   142
(* [| H |- p=>Fls;  H |- p;  q: prop |] ==> H |- q *)
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   143
bind_thm ("thms_notE", (thms_MP RS thms_FlsE));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   145
(*Soundness of the rules wrt truth-table semantics*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   146
goalw PropLog.thy [logcon_def] "!!H. H |- p ==> H |= p";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   147
by (etac thms.induct 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   148
by (fast_tac (ZF_cs addSDs [is_true_Imp RS iffD1 RS mp]) 5);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   149
by (ALLGOALS (asm_simp_tac prop_ss));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   150
qed "soundness";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   151
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   152
(*** Towards the completeness proof ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   153
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   154
val [premf,premq] = goal PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   155
    "[| H |- p=>Fls; q: prop |] ==> H |- p=>q";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
by (rtac (premf RS thms_in_pl RS ImpE) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   157
by (rtac deduction 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   158
by (rtac (premf RS weaken_left_cons RS thms_notE) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   159
by (REPEAT (ares_tac [premq, consI1, thms.H] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   160
qed "Fls_Imp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   161
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   162
val [premp,premq] = goal PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   163
    "[| H |- p;  H |- q=>Fls |] ==> H |- (p=>q)=>Fls";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
by (cut_facts_tac ([premp,premq] RL [thms_in_pl]) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   165
by (etac ImpE 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
by (rtac deduction 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   167
by (rtac (premq RS weaken_left_cons RS thms_MP) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   168
by (rtac (consI1 RS thms.H RS thms_MP) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   169
by (rtac (premp RS weaken_left_cons) 2);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   170
by (REPEAT (ares_tac prop.intrs 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   171
qed "Imp_Fls";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   172
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
(*Typical example of strengthening the induction formula*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   174
val [major] = goal PropLog.thy 
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
    "p: prop ==> hyps(p,t) |- if(is_true(p,t), p, p=>Fls)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
by (rtac (expand_if RS iffD2) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   177
by (rtac (major RS prop.induct) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   178
by (ALLGOALS (asm_simp_tac (prop_ss addsimps [thms_I, thms.H])));
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 7
diff changeset
   179
by (safe_tac (ZF_cs addSEs [Fls_Imp RS weaken_left_Un1, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   180
                            Fls_Imp RS weaken_left_Un2]));
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 7
diff changeset
   181
by (ALLGOALS (fast_tac (ZF_cs addIs [weaken_left_Un1, weaken_left_Un2, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   182
                                     weaken_right, Imp_Fls])));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   183
qed "hyps_thms_if";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   184
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   185
(*Key lemma for completeness; yields a set of assumptions satisfying p*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   186
val [premp,sat] = goalw PropLog.thy [logcon_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   187
    "[| p: prop;  0 |= p |] ==> hyps(p,t) |- p";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   188
by (rtac (sat RS spec RS mp RS if_P RS subst) 1 THEN
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   189
    rtac (premp RS hyps_thms_if) 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   190
by (fast_tac ZF_cs 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   191
qed "logcon_thms_p";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   192
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   193
(*For proving certain theorems in our new propositional logic*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   194
val thms_cs = 
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   195
    ZF_cs addSIs (prop.intrs @ [deduction])
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   196
          addIs [thms_in_pl, thms.H, thms.H RS thms_MP];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   197
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   198
(*The excluded middle in the form of an elimination rule*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   199
val prems = goal PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   200
    "[| p: prop;  q: prop |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   201
by (rtac (deduction RS deduction) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   202
by (rtac (thms.DN RS thms_MP) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   203
by (ALLGOALS (best_tac (thms_cs addSIs prems)));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   204
qed "thms_excluded_middle";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   205
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   206
(*Hard to prove directly because it requires cuts*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   207
val prems = goal PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   208
    "[| cons(p,H) |- q;  cons(p=>Fls,H) |- q;  p: prop |] ==> H |- q";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   209
by (rtac (thms_excluded_middle RS thms_MP RS thms_MP) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   210
by (REPEAT (resolve_tac (prems@prop.intrs@[deduction,thms_in_pl]) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   211
qed "thms_excluded_middle_rule";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   212
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   213
(*** Completeness -- lemmas for reducing the set of assumptions ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   214
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   215
(*For the case hyps(p,t)-cons(#v,Y) |- p;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
  we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   217
val [major] = goal PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   218
    "p: prop ==> hyps(p, t-{v}) <= cons(#v=>Fls, hyps(p,t)-{#v})";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   219
by (rtac (major RS prop.induct) 1);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   220
by (simp_tac prop_ss 1);
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   221
by (asm_simp_tac (prop_ss setloop (split_tac [expand_if])) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   222
by (fast_tac (ZF_cs addSEs prop.free_SEs) 1);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   223
by (asm_simp_tac prop_ss 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   224
by (fast_tac ZF_cs 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   225
qed "hyps_Diff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   226
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   227
(*For the case hyps(p,t)-cons(#v => Fls,Y) |- p;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   228
  we also have hyps(p,t)-{#v=>Fls} <= hyps(p, cons(v,t)) *)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   229
val [major] = goal PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   230
    "p: prop ==> hyps(p, cons(v,t)) <= cons(#v, hyps(p,t)-{#v=>Fls})";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   231
by (rtac (major RS prop.induct) 1);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   232
by (simp_tac prop_ss 1);
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   233
by (asm_simp_tac (prop_ss setloop (split_tac [expand_if])) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   234
by (fast_tac (ZF_cs addSEs prop.free_SEs) 1);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   235
by (asm_simp_tac prop_ss 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   236
by (fast_tac ZF_cs 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   237
qed "hyps_cons";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   238
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   239
(** Two lemmas for use with weaken_left **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   240
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   241
goal ZF.thy "B-C <= cons(a, B-cons(a,C))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   242
by (fast_tac ZF_cs 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   243
qed "cons_Diff_same";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   244
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   245
goal ZF.thy "cons(a, B-{c}) - D <= cons(a, B-cons(c,D))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   246
by (fast_tac ZF_cs 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   247
qed "cons_Diff_subset2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   248
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   249
(*The set hyps(p,t) is finite, and elements have the form #v or #v=>Fls;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   250
 could probably prove the stronger hyps(p,t) : Fin(hyps(p,0) Un hyps(p,nat))*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   251
val [major] = goal PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   252
    "p: prop ==> hyps(p,t) : Fin(UN v:nat. {#v, #v=>Fls})";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   253
by (rtac (major RS prop.induct) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   254
by (asm_simp_tac (prop_ss addsimps (Fin.intrs @ [UN_I, cons_iff])
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   255
                  setloop (split_tac [expand_if])) 2);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   256
by (ALLGOALS (asm_simp_tac (prop_ss addsimps [Un_0, Fin.emptyI, Fin_UnI])));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   257
qed "hyps_finite";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   258
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   259
val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   260
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   261
(*Induction on the finite set of assumptions hyps(p,t0).
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   262
  We may repeatedly subtract assumptions until none are left!*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   263
val [premp,sat] = goal PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   264
    "[| p: prop;  0 |= p |] ==> ALL t. hyps(p,t) - hyps(p,t0) |- p";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   265
by (rtac (premp RS hyps_finite RS Fin_induct) 1);
501
9c724f7085f9 ZF/ex/PropLog/sat_XXX: renamed logcon_XXX, since the relation is logical
lcp
parents: 477
diff changeset
   266
by (simp_tac (prop_ss addsimps [premp, sat, logcon_thms_p, Diff_0]) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   267
by (safe_tac ZF_cs);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   268
(*Case hyps(p,t)-cons(#v,Y) |- p *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   269
by (rtac thms_excluded_middle_rule 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   270
by (etac prop.Var_I 3);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   271
by (rtac (cons_Diff_same RS weaken_left) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   272
by (etac spec 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   273
by (rtac (cons_Diff_subset2 RS weaken_left) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   274
by (rtac (premp RS hyps_Diff RS Diff_weaken_left) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   275
by (etac spec 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   276
(*Case hyps(p,t)-cons(#v => Fls,Y) |- p *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   277
by (rtac thms_excluded_middle_rule 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   278
by (etac prop.Var_I 3);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   279
by (rtac (cons_Diff_same RS weaken_left) 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   280
by (etac spec 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   281
by (rtac (cons_Diff_subset2 RS weaken_left) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   282
by (rtac (premp RS hyps_cons RS Diff_weaken_left) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   283
by (etac spec 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   284
qed "completeness_0_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   285
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   286
(*The base case for completeness*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   287
val [premp,sat] = goal PropLog.thy "[| p: prop;  0 |= p |] ==> 0 |- p";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   288
by (rtac (Diff_cancel RS subst) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   289
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
   290
qed "completeness_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   291
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   292
(*A semantic analogue of the Deduction Theorem*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   293
goalw PropLog.thy [logcon_def] "!!H p q. [| cons(p,H) |= q |] ==> H |= p=>q";
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
   294
by (simp_tac prop_ss 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   295
by (fast_tac ZF_cs 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   296
qed "logcon_Imp";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   297
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   298
goal PropLog.thy "!!H. H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   299
by (etac Fin_induct 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   300
by (safe_tac (ZF_cs addSIs [completeness_0]));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   301
by (rtac (weaken_left_cons RS thms_MP) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   302
by (fast_tac (ZF_cs addSIs (logcon_Imp::prop.intrs)) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   303
by (fast_tac thms_cs 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   304
qed "completeness_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   305
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   306
val completeness = completeness_lemma RS bspec RS mp;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   307
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
   308
val [finite] = goal PropLog.thy "H: Fin(prop) ==> H |- p <-> H |= p & p:prop";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   309
by (fast_tac (ZF_cs addSEs [soundness, finite RS completeness, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   310
                            thms_in_pl]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   311
qed "thms_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   312
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   313
writeln"Reached end of file.";