src/ZF/Induct/PropLog.thy
changeset 60770 240563fbf41d
parent 58871 c399ae4b836f
child 61798 27f3c10b0b50
equal deleted inserted replaced
60769:cf7f3465eaf1 60770:240563fbf41d
     1 (*  Title:      ZF/Induct/PropLog.thy
     1 (*  Title:      ZF/Induct/PropLog.thy
     2     Author:     Tobias Nipkow & Lawrence C Paulson
     2     Author:     Tobias Nipkow & Lawrence C Paulson
     3     Copyright   1993  University of Cambridge
     3     Copyright   1993  University of Cambridge
     4 *)
     4 *)
     5 
     5 
     6 section {* Meta-theory of propositional logic *}
     6 section \<open>Meta-theory of propositional logic\<close>
     7 
     7 
     8 theory PropLog imports Main begin
     8 theory PropLog imports Main begin
     9 
     9 
    10 text {*
    10 text \<open>
    11   Datatype definition of propositional logic formulae and inductive
    11   Datatype definition of propositional logic formulae and inductive
    12   definition of the propositional tautologies.
    12   definition of the propositional tautologies.
    13 
    13 
    14   Inductive definition of propositional logic.  Soundness and
    14   Inductive definition of propositional logic.  Soundness and
    15   completeness w.r.t.\ truth-tables.
    15   completeness w.r.t.\ truth-tables.
    16 
    16 
    17   Prove: If @{text "H |= p"} then @{text "G |= p"} where @{text "G \<in>
    17   Prove: If @{text "H |= p"} then @{text "G |= p"} where @{text "G \<in>
    18   Fin(H)"}
    18   Fin(H)"}
    19 *}
    19 \<close>
    20 
    20 
    21 
    21 
    22 subsection {* The datatype of propositions *}
    22 subsection \<open>The datatype of propositions\<close>
    23 
    23 
    24 consts
    24 consts
    25   propn :: i
    25   propn :: i
    26 
    26 
    27 datatype propn =
    27 datatype propn =
    28     Fls
    28     Fls
    29   | Var ("n \<in> nat")    ("#_" [100] 100)
    29   | Var ("n \<in> nat")    ("#_" [100] 100)
    30   | Imp ("p \<in> propn", "q \<in> propn")    (infixr "=>" 90)
    30   | Imp ("p \<in> propn", "q \<in> propn")    (infixr "=>" 90)
    31 
    31 
    32 
    32 
    33 subsection {* The proof system *}
    33 subsection \<open>The proof system\<close>
    34 
    34 
    35 consts thms     :: "i => i"
    35 consts thms     :: "i => i"
    36 
    36 
    37 abbreviation
    37 abbreviation
    38   thms_syntax :: "[i,i] => o"    (infixl "|-" 50)
    38   thms_syntax :: "[i,i] => o"    (infixl "|-" 50)
    50   type_intros "propn.intros"
    50   type_intros "propn.intros"
    51 
    51 
    52 declare propn.intros [simp]
    52 declare propn.intros [simp]
    53 
    53 
    54 
    54 
    55 subsection {* The semantics *}
    55 subsection \<open>The semantics\<close>
    56 
    56 
    57 subsubsection {* Semantics of propositional logic. *}
    57 subsubsection \<open>Semantics of propositional logic.\<close>
    58 
    58 
    59 consts
    59 consts
    60   is_true_fun :: "[i,i] => i"
    60   is_true_fun :: "[i,i] => i"
    61 primrec
    61 primrec
    62   "is_true_fun(Fls, t) = 0"
    62   "is_true_fun(Fls, t) = 0"
    64   "is_true_fun(p=>q, t) = (if is_true_fun(p,t) = 1 then is_true_fun(q,t) else 1)"
    64   "is_true_fun(p=>q, t) = (if is_true_fun(p,t) = 1 then is_true_fun(q,t) else 1)"
    65 
    65 
    66 definition
    66 definition
    67   is_true :: "[i,i] => o"  where
    67   is_true :: "[i,i] => o"  where
    68   "is_true(p,t) == is_true_fun(p,t) = 1"
    68   "is_true(p,t) == is_true_fun(p,t) = 1"
    69   -- {* this definition is required since predicates can't be recursive *}
    69   -- \<open>this definition is required since predicates can't be recursive\<close>
    70 
    70 
    71 lemma is_true_Fls [simp]: "is_true(Fls,t) \<longleftrightarrow> False"
    71 lemma is_true_Fls [simp]: "is_true(Fls,t) \<longleftrightarrow> False"
    72   by (simp add: is_true_def)
    72   by (simp add: is_true_def)
    73 
    73 
    74 lemma is_true_Var [simp]: "is_true(#v,t) \<longleftrightarrow> v \<in> t"
    74 lemma is_true_Var [simp]: "is_true(#v,t) \<longleftrightarrow> v \<in> t"
    76 
    76 
    77 lemma is_true_Imp [simp]: "is_true(p=>q,t) \<longleftrightarrow> (is_true(p,t)\<longrightarrow>is_true(q,t))"
    77 lemma is_true_Imp [simp]: "is_true(p=>q,t) \<longleftrightarrow> (is_true(p,t)\<longrightarrow>is_true(q,t))"
    78   by (simp add: is_true_def)
    78   by (simp add: is_true_def)
    79 
    79 
    80 
    80 
    81 subsubsection {* Logical consequence *}
    81 subsubsection \<open>Logical consequence\<close>
    82 
    82 
    83 text {*
    83 text \<open>
    84   For every valuation, if all elements of @{text H} are true then so
    84   For every valuation, if all elements of @{text H} are true then so
    85   is @{text p}.
    85   is @{text p}.
    86 *}
    86 \<close>
    87 
    87 
    88 definition
    88 definition
    89   logcon :: "[i,i] => o"    (infixl "|=" 50)  where
    89   logcon :: "[i,i] => o"    (infixl "|=" 50)  where
    90   "H |= p == \<forall>t. (\<forall>q \<in> H. is_true(q,t)) \<longrightarrow> is_true(p,t)"
    90   "H |= p == \<forall>t. (\<forall>q \<in> H. is_true(q,t)) \<longrightarrow> is_true(p,t)"
    91 
    91 
    92 
    92 
    93 text {*
    93 text \<open>
    94   A finite set of hypotheses from @{text t} and the @{text Var}s in
    94   A finite set of hypotheses from @{text t} and the @{text Var}s in
    95   @{text p}.
    95   @{text p}.
    96 *}
    96 \<close>
    97 
    97 
    98 consts
    98 consts
    99   hyps :: "[i,i] => i"
    99   hyps :: "[i,i] => i"
   100 primrec
   100 primrec
   101   "hyps(Fls, t) = 0"
   101   "hyps(Fls, t) = 0"
   102   "hyps(Var(v), t) = (if v \<in> t then {#v} else {#v=>Fls})"
   102   "hyps(Var(v), t) = (if v \<in> t then {#v} else {#v=>Fls})"
   103   "hyps(p=>q, t) = hyps(p,t) \<union> hyps(q,t)"
   103   "hyps(p=>q, t) = hyps(p,t) \<union> hyps(q,t)"
   104 
   104 
   105 
   105 
   106 
   106 
   107 subsection {* Proof theory of propositional logic *}
   107 subsection \<open>Proof theory of propositional logic\<close>
   108 
   108 
   109 lemma thms_mono: "G \<subseteq> H ==> thms(G) \<subseteq> thms(H)"
   109 lemma thms_mono: "G \<subseteq> H ==> thms(G) \<subseteq> thms(H)"
   110   apply (unfold thms.defs)
   110   apply (unfold thms.defs)
   111   apply (rule lfp_mono)
   111   apply (rule lfp_mono)
   112     apply (rule thms.bnd_mono)+
   112     apply (rule thms.bnd_mono)+
   116 lemmas thms_in_pl = thms.dom_subset [THEN subsetD]
   116 lemmas thms_in_pl = thms.dom_subset [THEN subsetD]
   117 
   117 
   118 inductive_cases ImpE: "p=>q \<in> propn"
   118 inductive_cases ImpE: "p=>q \<in> propn"
   119 
   119 
   120 lemma thms_MP: "[| H |- p=>q;  H |- p |] ==> H |- q"
   120 lemma thms_MP: "[| H |- p=>q;  H |- p |] ==> H |- q"
   121   -- {* Stronger Modus Ponens rule: no typechecking! *}
   121   -- \<open>Stronger Modus Ponens rule: no typechecking!\<close>
   122   apply (rule thms.MP)
   122   apply (rule thms.MP)
   123      apply (erule asm_rl thms_in_pl thms_in_pl [THEN ImpE])+
   123      apply (erule asm_rl thms_in_pl thms_in_pl [THEN ImpE])+
   124   done
   124   done
   125 
   125 
   126 lemma thms_I: "p \<in> propn ==> H |- p=>p"
   126 lemma thms_I: "p \<in> propn ==> H |- p=>p"
   127   -- {*Rule is called @{text I} for Identity Combinator, not for Introduction. *}
   127   -- \<open>Rule is called @{text I} for Identity Combinator, not for Introduction.\<close>
   128   apply (rule thms.S [THEN thms_MP, THEN thms_MP])
   128   apply (rule thms.S [THEN thms_MP, THEN thms_MP])
   129       apply (rule_tac [5] thms.K)
   129       apply (rule_tac [5] thms.K)
   130        apply (rule_tac [4] thms.K)
   130        apply (rule_tac [4] thms.K)
   131          apply simp_all
   131          apply simp_all
   132   done
   132   done
   133 
   133 
   134 
   134 
   135 subsubsection {* Weakening, left and right *}
   135 subsubsection \<open>Weakening, left and right\<close>
   136 
   136 
   137 lemma weaken_left: "[| G \<subseteq> H;  G|-p |] ==> H|-p"
   137 lemma weaken_left: "[| G \<subseteq> H;  G|-p |] ==> H|-p"
   138   -- {* Order of premises is convenient with @{text THEN} *}
   138   -- \<open>Order of premises is convenient with @{text THEN}\<close>
   139   by (erule thms_mono [THEN subsetD])
   139   by (erule thms_mono [THEN subsetD])
   140 
   140 
   141 lemma weaken_left_cons: "H |- p ==> cons(a,H) |- p"
   141 lemma weaken_left_cons: "H |- p ==> cons(a,H) |- p"
   142   by (erule subset_consI [THEN weaken_left])
   142   by (erule subset_consI [THEN weaken_left])
   143 
   143 
   146 
   146 
   147 lemma weaken_right: "[| H |- q;  p \<in> propn |] ==> H |- p=>q"
   147 lemma weaken_right: "[| H |- q;  p \<in> propn |] ==> H |- p=>q"
   148   by (simp_all add: thms.K [THEN thms_MP] thms_in_pl)
   148   by (simp_all add: thms.K [THEN thms_MP] thms_in_pl)
   149 
   149 
   150 
   150 
   151 subsubsection {* The deduction theorem *}
   151 subsubsection \<open>The deduction theorem\<close>
   152 
   152 
   153 theorem deduction: "[| cons(p,H) |- q;  p \<in> propn |] ==>  H |- p=>q"
   153 theorem deduction: "[| cons(p,H) |- q;  p \<in> propn |] ==>  H |- p=>q"
   154   apply (erule thms.induct)
   154   apply (erule thms.induct)
   155       apply (blast intro: thms_I thms.H [THEN weaken_right])
   155       apply (blast intro: thms_I thms.H [THEN weaken_right])
   156      apply (blast intro: thms.K [THEN weaken_right])
   156      apply (blast intro: thms.K [THEN weaken_right])
   158    apply (blast intro: thms.DN [THEN weaken_right])
   158    apply (blast intro: thms.DN [THEN weaken_right])
   159   apply (blast intro: thms.S [THEN thms_MP [THEN thms_MP]])
   159   apply (blast intro: thms.S [THEN thms_MP [THEN thms_MP]])
   160   done
   160   done
   161 
   161 
   162 
   162 
   163 subsubsection {* The cut rule *}
   163 subsubsection \<open>The cut rule\<close>
   164 
   164 
   165 lemma cut: "[| H|-p;  cons(p,H) |- q |] ==>  H |- q"
   165 lemma cut: "[| H|-p;  cons(p,H) |- q |] ==>  H |- q"
   166   apply (rule deduction [THEN thms_MP])
   166   apply (rule deduction [THEN thms_MP])
   167     apply (simp_all add: thms_in_pl)
   167     apply (simp_all add: thms_in_pl)
   168   done
   168   done
   175 
   175 
   176 lemma thms_notE: "[| H |- p=>Fls;  H |- p;  q \<in> propn |] ==> H |- q"
   176 lemma thms_notE: "[| H |- p=>Fls;  H |- p;  q \<in> propn |] ==> H |- q"
   177   by (erule thms_MP [THEN thms_FlsE])
   177   by (erule thms_MP [THEN thms_FlsE])
   178 
   178 
   179 
   179 
   180 subsubsection {* Soundness of the rules wrt truth-table semantics *}
   180 subsubsection \<open>Soundness of the rules wrt truth-table semantics\<close>
   181 
   181 
   182 theorem soundness: "H |- p ==> H |= p"
   182 theorem soundness: "H |- p ==> H |= p"
   183   apply (unfold logcon_def)
   183   apply (unfold logcon_def)
   184   apply (induct set: thms)
   184   apply (induct set: thms)
   185       apply auto
   185       apply auto
   186   done
   186   done
   187 
   187 
   188 
   188 
   189 subsection {* Completeness *}
   189 subsection \<open>Completeness\<close>
   190 
   190 
   191 subsubsection {* Towards the completeness proof *}
   191 subsubsection \<open>Towards the completeness proof\<close>
   192 
   192 
   193 lemma Fls_Imp: "[| H |- p=>Fls; q \<in> propn |] ==> H |- p=>q"
   193 lemma Fls_Imp: "[| H |- p=>Fls; q \<in> propn |] ==> H |- p=>q"
   194   apply (frule thms_in_pl)
   194   apply (frule thms_in_pl)
   195   apply (rule deduction)
   195   apply (rule deduction)
   196    apply (rule weaken_left_cons [THEN thms_notE])
   196    apply (rule weaken_left_cons [THEN thms_notE])
   206     apply (blast intro: weaken_left_cons elim: ImpE)+
   206     apply (blast intro: weaken_left_cons elim: ImpE)+
   207   done
   207   done
   208 
   208 
   209 lemma hyps_thms_if:
   209 lemma hyps_thms_if:
   210     "p \<in> propn ==> hyps(p,t) |- (if is_true(p,t) then p else p=>Fls)"
   210     "p \<in> propn ==> hyps(p,t) |- (if is_true(p,t) then p else p=>Fls)"
   211   -- {* Typical example of strengthening the induction statement. *}
   211   -- \<open>Typical example of strengthening the induction statement.\<close>
   212   apply simp
   212   apply simp
   213   apply (induct_tac p)
   213   apply (induct_tac p)
   214     apply (simp_all add: thms_I thms.H)
   214     apply (simp_all add: thms_I thms.H)
   215   apply (safe elim!: Fls_Imp [THEN weaken_left_Un1] Fls_Imp [THEN weaken_left_Un2])
   215   apply (safe elim!: Fls_Imp [THEN weaken_left_Un1] Fls_Imp [THEN weaken_left_Un2])
   216   apply (blast intro: weaken_left_Un1 weaken_left_Un2 weaken_right Imp_Fls)+
   216   apply (blast intro: weaken_left_Un1 weaken_left_Un2 weaken_right Imp_Fls)+
   217   done
   217   done
   218 
   218 
   219 lemma logcon_thms_p: "[| p \<in> propn;  0 |= p |] ==> hyps(p,t) |- p"
   219 lemma logcon_thms_p: "[| p \<in> propn;  0 |= p |] ==> hyps(p,t) |- p"
   220   -- {* Key lemma for completeness; yields a set of assumptions satisfying @{text p} *}
   220   -- \<open>Key lemma for completeness; yields a set of assumptions satisfying @{text p}\<close>
   221   apply (drule hyps_thms_if)
   221   apply (drule hyps_thms_if)
   222   apply (simp add: logcon_def)
   222   apply (simp add: logcon_def)
   223   done
   223   done
   224 
   224 
   225 text {*
   225 text \<open>
   226   For proving certain theorems in our new propositional logic.
   226   For proving certain theorems in our new propositional logic.
   227 *}
   227 \<close>
   228 
   228 
   229 lemmas propn_SIs = propn.intros deduction
   229 lemmas propn_SIs = propn.intros deduction
   230   and propn_Is = thms_in_pl thms.H thms.H [THEN thms_MP]
   230   and propn_Is = thms_in_pl thms.H thms.H [THEN thms_MP]
   231 
   231 
   232 text {*
   232 text \<open>
   233   The excluded middle in the form of an elimination rule.
   233   The excluded middle in the form of an elimination rule.
   234 *}
   234 \<close>
   235 
   235 
   236 lemma thms_excluded_middle:
   236 lemma thms_excluded_middle:
   237     "[| p \<in> propn;  q \<in> propn |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q"
   237     "[| p \<in> propn;  q \<in> propn |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q"
   238   apply (rule deduction [THEN deduction])
   238   apply (rule deduction [THEN deduction])
   239     apply (rule thms.DN [THEN thms_MP])
   239     apply (rule thms.DN [THEN thms_MP])
   240      apply (best intro!: propn_SIs intro: propn_Is)+
   240      apply (best intro!: propn_SIs intro: propn_Is)+
   241   done
   241   done
   242 
   242 
   243 lemma thms_excluded_middle_rule:
   243 lemma thms_excluded_middle_rule:
   244   "[| cons(p,H) |- q;  cons(p=>Fls,H) |- q;  p \<in> propn |] ==> H |- q"
   244   "[| cons(p,H) |- q;  cons(p=>Fls,H) |- q;  p \<in> propn |] ==> H |- q"
   245   -- {* Hard to prove directly because it requires cuts *}
   245   -- \<open>Hard to prove directly because it requires cuts\<close>
   246   apply (rule thms_excluded_middle [THEN thms_MP, THEN thms_MP])
   246   apply (rule thms_excluded_middle [THEN thms_MP, THEN thms_MP])
   247      apply (blast intro!: propn_SIs intro: propn_Is)+
   247      apply (blast intro!: propn_SIs intro: propn_Is)+
   248   done
   248   done
   249 
   249 
   250 
   250 
   251 subsubsection {* Completeness -- lemmas for reducing the set of assumptions *}
   251 subsubsection \<open>Completeness -- lemmas for reducing the set of assumptions\<close>
   252 
   252 
   253 text {*
   253 text \<open>
   254   For the case @{prop "hyps(p,t)-cons(#v,Y) |- p"} we also have @{prop
   254   For the case @{prop "hyps(p,t)-cons(#v,Y) |- p"} we also have @{prop
   255   "hyps(p,t)-{#v} \<subseteq> hyps(p, t-{v})"}.
   255   "hyps(p,t)-{#v} \<subseteq> hyps(p, t-{v})"}.
   256 *}
   256 \<close>
   257 
   257 
   258 lemma hyps_Diff:
   258 lemma hyps_Diff:
   259     "p \<in> propn ==> hyps(p, t-{v}) \<subseteq> cons(#v=>Fls, hyps(p,t)-{#v})"
   259     "p \<in> propn ==> hyps(p, t-{v}) \<subseteq> cons(#v=>Fls, hyps(p,t)-{#v})"
   260   by (induct set: propn) auto
   260   by (induct set: propn) auto
   261 
   261 
   262 text {*
   262 text \<open>
   263   For the case @{prop "hyps(p,t)-cons(#v => Fls,Y) |- p"} we also have
   263   For the case @{prop "hyps(p,t)-cons(#v => Fls,Y) |- p"} we also have
   264   @{prop "hyps(p,t)-{#v=>Fls} \<subseteq> hyps(p, cons(v,t))"}.
   264   @{prop "hyps(p,t)-{#v=>Fls} \<subseteq> hyps(p, cons(v,t))"}.
   265 *}
   265 \<close>
   266 
   266 
   267 lemma hyps_cons:
   267 lemma hyps_cons:
   268     "p \<in> propn ==> hyps(p, cons(v,t)) \<subseteq> cons(#v, hyps(p,t)-{#v=>Fls})"
   268     "p \<in> propn ==> hyps(p, cons(v,t)) \<subseteq> cons(#v, hyps(p,t)-{#v=>Fls})"
   269   by (induct set: propn) auto
   269   by (induct set: propn) auto
   270 
   270 
   271 text {* Two lemmas for use with @{text weaken_left} *}
   271 text \<open>Two lemmas for use with @{text weaken_left}\<close>
   272 
   272 
   273 lemma cons_Diff_same: "B-C \<subseteq> cons(a, B-cons(a,C))"
   273 lemma cons_Diff_same: "B-C \<subseteq> cons(a, B-cons(a,C))"
   274   by blast
   274   by blast
   275 
   275 
   276 lemma cons_Diff_subset2: "cons(a, B-{c}) - D \<subseteq> cons(a, B-cons(c,D))"
   276 lemma cons_Diff_subset2: "cons(a, B-{c}) - D \<subseteq> cons(a, B-cons(c,D))"
   277   by blast
   277   by blast
   278 
   278 
   279 text {*
   279 text \<open>
   280   The set @{term "hyps(p,t)"} is finite, and elements have the form
   280   The set @{term "hyps(p,t)"} is finite, and elements have the form
   281   @{term "#v"} or @{term "#v=>Fls"}; could probably prove the stronger
   281   @{term "#v"} or @{term "#v=>Fls"}; could probably prove the stronger
   282   @{prop "hyps(p,t) \<in> Fin(hyps(p,0) \<union> hyps(p,nat))"}.
   282   @{prop "hyps(p,t) \<in> Fin(hyps(p,0) \<union> hyps(p,nat))"}.
   283 *}
   283 \<close>
   284 
   284 
   285 lemma hyps_finite: "p \<in> propn ==> hyps(p,t) \<in> Fin(\<Union>v \<in> nat. {#v, #v=>Fls})"
   285 lemma hyps_finite: "p \<in> propn ==> hyps(p,t) \<in> Fin(\<Union>v \<in> nat. {#v, #v=>Fls})"
   286   by (induct set: propn) auto
   286   by (induct set: propn) auto
   287 
   287 
   288 lemmas Diff_weaken_left = Diff_mono [OF _ subset_refl, THEN weaken_left]
   288 lemmas Diff_weaken_left = Diff_mono [OF _ subset_refl, THEN weaken_left]
   289 
   289 
   290 text {*
   290 text \<open>
   291   Induction on the finite set of assumptions @{term "hyps(p,t0)"}.  We
   291   Induction on the finite set of assumptions @{term "hyps(p,t0)"}.  We
   292   may repeatedly subtract assumptions until none are left!
   292   may repeatedly subtract assumptions until none are left!
   293 *}
   293 \<close>
   294 
   294 
   295 lemma completeness_0_lemma [rule_format]:
   295 lemma completeness_0_lemma [rule_format]:
   296     "[| p \<in> propn;  0 |= p |] ==> \<forall>t. hyps(p,t) - hyps(p,t0) |- p"
   296     "[| p \<in> propn;  0 |= p |] ==> \<forall>t. hyps(p,t) - hyps(p,t0) |- p"
   297   apply (frule hyps_finite)
   297   apply (frule hyps_finite)
   298   apply (erule Fin_induct)
   298   apply (erule Fin_induct)
   299    apply (simp add: logcon_thms_p Diff_0)
   299    apply (simp add: logcon_thms_p Diff_0)
   300   txt {* inductive step *}
   300   txt \<open>inductive step\<close>
   301   apply safe
   301   apply safe
   302    txt {* Case @{prop "hyps(p,t)-cons(#v,Y) |- p"} *}
   302    txt \<open>Case @{prop "hyps(p,t)-cons(#v,Y) |- p"}\<close>
   303    apply (rule thms_excluded_middle_rule)
   303    apply (rule thms_excluded_middle_rule)
   304      apply (erule_tac [3] propn.intros)
   304      apply (erule_tac [3] propn.intros)
   305     apply (blast intro: cons_Diff_same [THEN weaken_left])
   305     apply (blast intro: cons_Diff_same [THEN weaken_left])
   306    apply (blast intro: cons_Diff_subset2 [THEN weaken_left]
   306    apply (blast intro: cons_Diff_subset2 [THEN weaken_left]
   307      hyps_Diff [THEN Diff_weaken_left])
   307      hyps_Diff [THEN Diff_weaken_left])
   308   txt {* Case @{prop "hyps(p,t)-cons(#v => Fls,Y) |- p"} *}
   308   txt \<open>Case @{prop "hyps(p,t)-cons(#v => Fls,Y) |- p"}\<close>
   309   apply (rule thms_excluded_middle_rule)
   309   apply (rule thms_excluded_middle_rule)
   310     apply (erule_tac [3] propn.intros)
   310     apply (erule_tac [3] propn.intros)
   311    apply (blast intro: cons_Diff_subset2 [THEN weaken_left]
   311    apply (blast intro: cons_Diff_subset2 [THEN weaken_left]
   312      hyps_cons [THEN Diff_weaken_left])
   312      hyps_cons [THEN Diff_weaken_left])
   313   apply (blast intro: cons_Diff_same [THEN weaken_left])
   313   apply (blast intro: cons_Diff_same [THEN weaken_left])
   314   done
   314   done
   315 
   315 
   316 
   316 
   317 subsubsection {* Completeness theorem *}
   317 subsubsection \<open>Completeness theorem\<close>
   318 
   318 
   319 lemma completeness_0: "[| p \<in> propn;  0 |= p |] ==> 0 |- p"
   319 lemma completeness_0: "[| p \<in> propn;  0 |= p |] ==> 0 |- p"
   320   -- {* The base case for completeness *}
   320   -- \<open>The base case for completeness\<close>
   321   apply (rule Diff_cancel [THEN subst])
   321   apply (rule Diff_cancel [THEN subst])
   322   apply (blast intro: completeness_0_lemma)
   322   apply (blast intro: completeness_0_lemma)
   323   done
   323   done
   324 
   324 
   325 lemma logcon_Imp: "[| cons(p,H) |= q |] ==> H |= p=>q"
   325 lemma logcon_Imp: "[| cons(p,H) |= q |] ==> H |= p=>q"
   326   -- {* A semantic analogue of the Deduction Theorem *}
   326   -- \<open>A semantic analogue of the Deduction Theorem\<close>
   327   by (simp add: logcon_def)
   327   by (simp add: logcon_def)
   328 
   328 
   329 lemma completeness:
   329 lemma completeness:
   330      "H \<in> Fin(propn) ==> p \<in> propn \<Longrightarrow> H |= p \<Longrightarrow> H |- p"
   330      "H \<in> Fin(propn) ==> p \<in> propn \<Longrightarrow> H |= p \<Longrightarrow> H |- p"
   331   apply (induct arbitrary: p set: Fin)
   331   apply (induct arbitrary: p set: Fin)