src/ZF/ex/PropLog.ML
changeset 0 a5a9c433f639
child 7 268f93ab3bc4
equal deleted inserted replaced
-1:000000000000 0:a5a9c433f639
       
     1 (*  Title: 	ZF/ex/prop-log.ML
       
     2     ID:         $Id$
       
     3     Author: 	Tobias Nipkow & Lawrence C Paulson
       
     4     Copyright   1992  University of Cambridge
       
     5 
       
     6 For ex/prop-log.thy.  Inductive definition of propositional logic.
       
     7 Soundness and completeness w.r.t. truth-tables.
       
     8 
       
     9 Prove: If H|=p then G|=p where G:Fin(H)
       
    10 *)
       
    11 
       
    12 open PropLog;
       
    13 
       
    14 (*** prop_rec -- by Vset recursion ***)
       
    15 
       
    16 val prop_congs = mk_typed_congs Prop.thy 
       
    17 		   [("b", "[i,i,i]=>i"), ("d", "[i,i,i,i]=>i")];
       
    18 
       
    19 (** conversion rules **)
       
    20 
       
    21 goal PropLog.thy "prop_rec(Fls,b,c,d) = b";
       
    22 by (rtac (prop_rec_def RS def_Vrec RS trans) 1);
       
    23 by (rewrite_goals_tac Prop.con_defs);
       
    24 by (SIMP_TAC rank_ss 1);
       
    25 val prop_rec_Fls = result();
       
    26 
       
    27 goal PropLog.thy "prop_rec(#v,b,c,d) = c(v)";
       
    28 by (rtac (prop_rec_def RS def_Vrec RS trans) 1);
       
    29 by (rewrite_goals_tac Prop.con_defs);
       
    30 by (SIMP_TAC (rank_ss addcongs prop_congs) 1);
       
    31 val prop_rec_Var = result();
       
    32 
       
    33 goal PropLog.thy "prop_rec(p=>q,b,c,d) = \
       
    34 \      d(p, q, prop_rec(p,b,c,d), prop_rec(q,b,c,d))";
       
    35 by (rtac (prop_rec_def RS def_Vrec RS trans) 1);
       
    36 by (rewrite_goals_tac Prop.con_defs);
       
    37 by (SIMP_TAC (rank_ss addcongs prop_congs) 1);
       
    38 val prop_rec_Imp = result();
       
    39 
       
    40 val prop_rec_ss = 
       
    41     arith_ss addrews [prop_rec_Fls, prop_rec_Var, prop_rec_Imp];
       
    42 
       
    43 (*** Semantics of propositional logic ***)
       
    44 
       
    45 (** The function is_true **)
       
    46 
       
    47 goalw PropLog.thy [is_true_def] "is_true(Fls,t) <-> False";
       
    48 by (SIMP_TAC (prop_rec_ss addrews [one_not_0 RS not_sym]) 1);
       
    49 val is_true_Fls = result();
       
    50 
       
    51 goalw PropLog.thy [is_true_def] "is_true(#v,t) <-> v:t";
       
    52 by (SIMP_TAC (prop_rec_ss addrews [one_not_0 RS not_sym] 
       
    53 	      addsplits [expand_if]) 1);
       
    54 val is_true_Var = result();
       
    55 
       
    56 goalw PropLog.thy [is_true_def]
       
    57     "is_true(p=>q,t) <-> (is_true(p,t)-->is_true(q,t))";
       
    58 by (SIMP_TAC (prop_rec_ss addsplits [expand_if]) 1);
       
    59 val is_true_Imp = result();
       
    60 
       
    61 (** The function hyps **)
       
    62 
       
    63 goalw PropLog.thy [hyps_def] "hyps(Fls,t) = 0";
       
    64 by (SIMP_TAC prop_rec_ss 1);
       
    65 val hyps_Fls = result();
       
    66 
       
    67 goalw PropLog.thy [hyps_def] "hyps(#v,t) = {if(v:t, #v, #v=>Fls)}";
       
    68 by (SIMP_TAC prop_rec_ss 1);
       
    69 val hyps_Var = result();
       
    70 
       
    71 goalw PropLog.thy [hyps_def] "hyps(p=>q,t) = hyps(p,t) Un hyps(q,t)";
       
    72 by (SIMP_TAC prop_rec_ss 1);
       
    73 val hyps_Imp = result();
       
    74 
       
    75 val prop_ss = prop_rec_ss 
       
    76     addcongs Prop.congs
       
    77     addcongs (mk_congs PropLog.thy ["Fin", "thms", "op |=","is_true","hyps"])
       
    78     addrews Prop.intrs
       
    79     addrews [is_true_Fls, is_true_Var, is_true_Imp,
       
    80 	     hyps_Fls, hyps_Var, hyps_Imp];
       
    81 
       
    82 (*** Proof theory of propositional logic ***)
       
    83 
       
    84 structure PropThms = Inductive_Fun
       
    85  (val thy = PropLog.thy;
       
    86   val rec_doms = [("thms","prop")];
       
    87   val sintrs = 
       
    88       ["[| p:H;  p:prop |] ==> H |- p",
       
    89        "[| p:prop;  q:prop |] ==> H |- p=>q=>p",
       
    90        "[| p:prop;  q:prop;  r:prop |] ==> H |- (p=>q=>r) => (p=>q) => p=>r",
       
    91        "p:prop ==> H |- ((p=>Fls) => Fls) => p",
       
    92        "[| H |- p=>q;  H |- p;  p:prop;  q:prop |] ==> H |- q"];
       
    93   val monos = [];
       
    94   val con_defs = [];
       
    95   val type_intrs = Prop.intrs;
       
    96   val type_elims = []);
       
    97 
       
    98 goalw PropThms.thy PropThms.defs "!!G H. G<=H ==> thms(G) <= thms(H)";
       
    99 by (rtac lfp_mono 1);
       
   100 by (REPEAT (rtac PropThms.bnd_mono 1));
       
   101 by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
       
   102 val thms_mono = result();
       
   103 
       
   104 val thms_in_pl = PropThms.dom_subset RS subsetD;
       
   105 
       
   106 val [thms_H, thms_K, thms_S, thms_DN, weak_thms_MP] = PropThms.intrs;
       
   107 
       
   108 (*Modus Ponens rule -- this stronger version avoids typecheck*)
       
   109 goal PropThms.thy "!!p q H. [| H |- p=>q;  H |- p |] ==> H |- q";
       
   110 by (rtac weak_thms_MP 1);
       
   111 by (REPEAT (eresolve_tac [asm_rl, thms_in_pl, thms_in_pl RS ImpE] 1));
       
   112 val thms_MP = result();
       
   113 
       
   114 (*Rule is called I for Identity Combinator, not for Introduction*)
       
   115 goal PropThms.thy "!!p H. p:prop ==> H |- p=>p";
       
   116 by (rtac (thms_S RS thms_MP RS thms_MP) 1);
       
   117 by (rtac thms_K 5);
       
   118 by (rtac thms_K 4);
       
   119 by (REPEAT (ares_tac [ImpI] 1));
       
   120 val thms_I = result();
       
   121 
       
   122 (** Weakening, left and right **)
       
   123 
       
   124 (* [| G<=H;  G|-p |] ==> H|-p   Order of premises is convenient with RS*)
       
   125 val weaken_left = standard (thms_mono RS subsetD);
       
   126 
       
   127 (* H |- p ==> cons(a,H) |- p *)
       
   128 val weaken_left_cons = subset_consI RS weaken_left;
       
   129 
       
   130 val weaken_left_Un1  = Un_upper1 RS weaken_left;
       
   131 val weaken_left_Un2  = Un_upper2 RS weaken_left;
       
   132 
       
   133 goal PropThms.thy "!!H p q. [| H |- q;  p:prop |] ==> H |- p=>q";
       
   134 by (rtac (thms_K RS thms_MP) 1);
       
   135 by (REPEAT (ares_tac [thms_in_pl] 1));
       
   136 val weaken_right = result();
       
   137 
       
   138 (*The deduction theorem*)
       
   139 goal PropThms.thy "!!p q H. [| cons(p,H) |- q;  p:prop |] ==>  H |- p=>q";
       
   140 by (etac PropThms.induct 1);
       
   141 by (fast_tac (ZF_cs addIs [thms_I, thms_H RS weaken_right]) 1);
       
   142 by (fast_tac (ZF_cs addIs [thms_K RS weaken_right]) 1);
       
   143 by (fast_tac (ZF_cs addIs [thms_S RS weaken_right]) 1);
       
   144 by (fast_tac (ZF_cs addIs [thms_DN RS weaken_right]) 1);
       
   145 by (fast_tac (ZF_cs addIs [thms_S RS thms_MP RS thms_MP]) 1);
       
   146 val deduction = result();
       
   147 
       
   148 
       
   149 (*The cut rule*)
       
   150 goal PropThms.thy "!!H p q. [| H|-p;  cons(p,H) |- q |] ==>  H |- q";
       
   151 by (rtac (deduction RS thms_MP) 1);
       
   152 by (REPEAT (ares_tac [thms_in_pl] 1));
       
   153 val cut = result();
       
   154 
       
   155 goal PropThms.thy "!!H p. [| H |- Fls; p:prop |] ==> H |- p";
       
   156 by (rtac (thms_DN RS thms_MP) 1);
       
   157 by (rtac weaken_right 2);
       
   158 by (REPEAT (ares_tac (Prop.intrs@[consI1]) 1));
       
   159 val thms_FlsE = result();
       
   160 
       
   161 (* [| H |- p=>Fls;  H |- p;  q: prop |] ==> H |- q *)
       
   162 val thms_notE = standard (thms_MP RS thms_FlsE);
       
   163 
       
   164 (*Soundness of the rules wrt truth-table semantics*)
       
   165 val [major] = goalw PropThms.thy [sat_def] "H |- p ==> H |= p";
       
   166 by (rtac (major RS PropThms.induct) 1);
       
   167 by (fast_tac (ZF_cs addSDs [is_true_Imp RS iffD1 RS mp]) 5);
       
   168 by (ALLGOALS (SIMP_TAC prop_ss));
       
   169 val soundness = result();
       
   170 
       
   171 (*** Towards the completeness proof ***)
       
   172 
       
   173 val [premf,premq] = goal PropThms.thy
       
   174     "[| H |- p=>Fls; q: prop |] ==> H |- p=>q";
       
   175 by (rtac (premf RS thms_in_pl RS ImpE) 1);
       
   176 by (rtac deduction 1);
       
   177 by (rtac (premf RS weaken_left_cons RS thms_notE) 1);
       
   178 by (REPEAT (ares_tac [premq, consI1, thms_H] 1));
       
   179 val Fls_Imp = result();
       
   180 
       
   181 val [premp,premq] = goal PropThms.thy
       
   182     "[| H |- p;  H |- q=>Fls |] ==> H |- (p=>q)=>Fls";
       
   183 by (cut_facts_tac ([premp,premq] RL [thms_in_pl]) 1);
       
   184 by (etac ImpE 1);
       
   185 by (rtac deduction 1);
       
   186 by (rtac (premq RS weaken_left_cons RS thms_MP) 1);
       
   187 by (rtac (consI1 RS thms_H RS thms_MP) 1);
       
   188 by (rtac (premp RS weaken_left_cons) 2);
       
   189 by (REPEAT (ares_tac Prop.intrs 1));
       
   190 val Imp_Fls = result();
       
   191 
       
   192 (*Typical example of strengthening the induction formula*)
       
   193 val [major] = goal PropThms.thy 
       
   194     "p: prop ==> hyps(p,t) |- if(is_true(p,t), p, p=>Fls)";
       
   195 by (rtac (expand_if RS iffD2) 1);
       
   196 by (rtac (major RS Prop.induct) 1);
       
   197 by (ALLGOALS (ASM_SIMP_TAC (prop_ss addrews [thms_I, thms_H])));
       
   198 by (fast_tac (ZF_cs addIs [weaken_left_Un1, weaken_left_Un2, 
       
   199 			   weaken_right, Imp_Fls]
       
   200                     addSEs [Fls_Imp]) 1);
       
   201 val hyps_thms_if = result();
       
   202 
       
   203 (*Key lemma for completeness; yields a set of assumptions satisfying p*)
       
   204 val [premp,sat] = goalw PropThms.thy [sat_def]
       
   205     "[| p: prop;  0 |= p |] ==> hyps(p,t) |- p";
       
   206 by (rtac (sat RS spec RS mp RS if_P RS subst) 1 THEN
       
   207     rtac (premp RS hyps_thms_if) 2);
       
   208 by (fast_tac ZF_cs 1);
       
   209 val sat_thms_p = result();
       
   210 
       
   211 (*For proving certain theorems in our new propositional logic*)
       
   212 val thms_cs = 
       
   213     ZF_cs addSIs [FlsI, VarI, ImpI, deduction]
       
   214           addIs [thms_in_pl, thms_H, thms_H RS thms_MP];
       
   215 
       
   216 (*The excluded middle in the form of an elimination rule*)
       
   217 val prems = goal PropThms.thy
       
   218     "[| p: prop;  q: prop |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q";
       
   219 by (rtac (deduction RS deduction) 1);
       
   220 by (rtac (thms_DN RS thms_MP) 1);
       
   221 by (ALLGOALS (best_tac (thms_cs addSIs prems)));
       
   222 val thms_excluded_middle = result();
       
   223 
       
   224 (*Hard to prove directly because it requires cuts*)
       
   225 val prems = goal PropThms.thy
       
   226     "[| cons(p,H) |- q;  cons(p=>Fls,H) |- q;  p: prop |] ==> H |- q";
       
   227 by (rtac (thms_excluded_middle RS thms_MP RS thms_MP) 1);
       
   228 by (REPEAT (resolve_tac (prems@Prop.intrs@[deduction,thms_in_pl]) 1));
       
   229 val thms_excluded_middle_rule = result();
       
   230 
       
   231 (*** Completeness -- lemmas for reducing the set of assumptions ***)
       
   232 
       
   233 (*For the case hyps(p,t)-cons(#v,Y) |- p;
       
   234   we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *)
       
   235 val [major] = goal PropThms.thy
       
   236     "p: prop ==> hyps(p, t-{v}) <= cons(#v=>Fls, hyps(p,t)-{#v})";
       
   237 by (rtac (major RS Prop.induct) 1);
       
   238 by (SIMP_TAC prop_ss 1);
       
   239 by (ASM_SIMP_TAC (prop_ss addsplits [expand_if]) 1);
       
   240 by (fast_tac (ZF_cs addSEs Prop.free_SEs) 1);
       
   241 by (ASM_SIMP_TAC prop_ss 1);
       
   242 by (fast_tac ZF_cs 1);
       
   243 val hyps_Diff = result();
       
   244 
       
   245 (*For the case hyps(p,t)-cons(#v => Fls,Y) |- p;
       
   246   we also have hyps(p,t)-{#v=>Fls} <= hyps(p, cons(v,t)) *)
       
   247 val [major] = goal PropThms.thy
       
   248     "p: prop ==> hyps(p, cons(v,t)) <= cons(#v, hyps(p,t)-{#v=>Fls})";
       
   249 by (rtac (major RS Prop.induct) 1);
       
   250 by (SIMP_TAC prop_ss 1);
       
   251 by (ASM_SIMP_TAC (prop_ss addsplits [expand_if]) 1);
       
   252 by (fast_tac (ZF_cs addSEs Prop.free_SEs) 1);
       
   253 by (ASM_SIMP_TAC prop_ss 1);
       
   254 by (fast_tac ZF_cs 1);
       
   255 val hyps_cons = result();
       
   256 
       
   257 (** Two lemmas for use with weaken_left **)
       
   258 
       
   259 goal ZF.thy "B-C <= cons(a, B-cons(a,C))";
       
   260 by (fast_tac ZF_cs 1);
       
   261 val cons_Diff_same = result();
       
   262 
       
   263 goal ZF.thy "cons(a, B-{c}) - D <= cons(a, B-cons(c,D))";
       
   264 by (fast_tac ZF_cs 1);
       
   265 val cons_Diff_subset2 = result();
       
   266 
       
   267 (*The set hyps(p,t) is finite, and elements have the form #v or #v=>Fls;
       
   268  could probably prove the stronger hyps(p,t) : Fin(hyps(p,0) Un hyps(p,nat))*)
       
   269 val [major] = goal PropThms.thy
       
   270     "p: prop ==> hyps(p,t) : Fin(UN v:nat. {#v, #v=>Fls})";
       
   271 by (rtac (major RS Prop.induct) 1);
       
   272 by (ASM_SIMP_TAC (prop_ss addrews [Fin_0I, Fin_consI, UN_I] 
       
   273 		  addsplits [expand_if]) 2);
       
   274 by (ALLGOALS (ASM_SIMP_TAC (prop_ss addrews [Un_0, Fin_0I, Fin_UnI])));
       
   275 val hyps_finite = result();
       
   276 
       
   277 val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left;
       
   278 
       
   279 (*Induction on the finite set of assumptions hyps(p,t0).
       
   280   We may repeatedly subtract assumptions until none are left!*)
       
   281 val [premp,sat] = goal PropThms.thy
       
   282     "[| p: prop;  0 |= p |] ==> ALL t. hyps(p,t) - hyps(p,t0) |- p";
       
   283 by (rtac (premp RS hyps_finite RS Fin_induct) 1);
       
   284 by (SIMP_TAC (prop_ss addrews [premp, sat, sat_thms_p, Diff_0]) 1);
       
   285 by (safe_tac ZF_cs);
       
   286 (*Case hyps(p,t)-cons(#v,Y) |- p *)
       
   287 by (rtac thms_excluded_middle_rule 1);
       
   288 by (etac VarI 3);
       
   289 by (rtac (cons_Diff_same RS weaken_left) 1);
       
   290 by (etac spec 1);
       
   291 by (rtac (cons_Diff_subset2 RS weaken_left) 1);
       
   292 by (rtac (premp RS hyps_Diff RS Diff_weaken_left) 1);
       
   293 by (etac spec 1);
       
   294 (*Case hyps(p,t)-cons(#v => Fls,Y) |- p *)
       
   295 by (rtac thms_excluded_middle_rule 1);
       
   296 by (etac VarI 3);
       
   297 by (rtac (cons_Diff_same RS weaken_left) 2);
       
   298 by (etac spec 2);
       
   299 by (rtac (cons_Diff_subset2 RS weaken_left) 1);
       
   300 by (rtac (premp RS hyps_cons RS Diff_weaken_left) 1);
       
   301 by (etac spec 1);
       
   302 val completeness_0_lemma = result();
       
   303 
       
   304 (*The base case for completeness*)
       
   305 val [premp,sat] = goal PropThms.thy "[| p: prop;  0 |= p |] ==> 0 |- p";
       
   306 by (rtac (Diff_cancel RS subst) 1);
       
   307 by (rtac (sat RS (premp RS completeness_0_lemma RS spec)) 1);
       
   308 val completeness_0 = result();
       
   309 
       
   310 (*A semantic analogue of the Deduction Theorem*)
       
   311 goalw PropThms.thy [sat_def] "!!H p q. [| cons(p,H) |= q |] ==> H |= p=>q";
       
   312 by (SIMP_TAC prop_ss 1);
       
   313 by (fast_tac ZF_cs 1);
       
   314 val sat_Imp = result();
       
   315 
       
   316 goal PropThms.thy "!!H. H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p";
       
   317 by (etac Fin_induct 1);
       
   318 by (safe_tac (ZF_cs addSIs [completeness_0]));
       
   319 by (rtac (weaken_left_cons RS thms_MP) 1);
       
   320 by (fast_tac (ZF_cs addSIs [sat_Imp,ImpI]) 1);
       
   321 by (fast_tac thms_cs 1);
       
   322 val completeness_lemma = result();
       
   323 
       
   324 val completeness = completeness_lemma RS bspec RS mp;
       
   325 
       
   326 val [finite] = goal PropThms.thy "H: Fin(prop) ==> H |- p <-> H |= p & p:prop";
       
   327 by (fast_tac (ZF_cs addSEs [soundness, finite RS completeness, 
       
   328 			    thms_in_pl]) 1);
       
   329 val thms_iff = result();
       
   330 
       
   331 writeln"Reached end of file.";
       
   332 
       
   333