X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fdemo%2Fpropositional_sequent_calculus.ma;h=42ae9682532ce8f65811f23944ec7e0f2a83cf0c;hb=e78cf74f8976cf0ca554f64baa9979d0423ee927;hp=36bb7e86308d07d2869fb0ad45d484d6bb9e8a1a;hpb=9eabe046c1182960de8cfdba96c5414224e3a61e;p=helm.git diff --git a/helm/software/matita/library/demo/propositional_sequent_calculus.ma b/helm/software/matita/library/demo/propositional_sequent_calculus.ma index 36bb7e863..42ae96825 100644 --- a/helm/software/matita/library/demo/propositional_sequent_calculus.ma +++ b/helm/software/matita/library/demo/propositional_sequent_calculus.ma @@ -181,6 +181,8 @@ axiom symm_andb: symmetric ? andb. axiom associative_andb: associative ? andb. axiom distributive_andb_orb: distributive ? andb orb. +lemma andb_true: ∀x.(true ∧ x) = x. intro; reflexivity. qed. + lemma and_of_list_permut: ∀i,f,l1,l2. eval i (and_of_list (l1 @ (f::l2))) = eval i (and_of_list (f :: l1 @ l2)). intros; @@ -242,11 +244,20 @@ theorem soundness: ∀F. derive F → is_tautology (formula_of_sequent F). lapply (H4 i); clear H4; rewrite > symm_orb in ⊢ (? ? (? ? %) ?); rewrite > distributive_orb_andb; + demodulate. + reflexivity. + (* autobatch paramodulation + by distributive_orb_andb,symm_orb,symm_orb, + Hletin, Hletin1,andb_true. + *) | simplify in H2 ⊢ %; intros; lapply (H2 i); clear H2; - autobatch + pump 100. pump 100. + demodulate. + reflexivity. + (* autobatch paramodulation by andb_assoc, Hletin. *) | simplify in H2 H4 ⊢ %; intros; lapply (H2 i); clear H2; @@ -256,11 +267,13 @@ theorem soundness: ∀F. derive F → is_tautology (formula_of_sequent F). rewrite > demorgan2; rewrite > symm_orb; rewrite > distributive_orb_andb; - autobatch paramodulation + demodulate. + reflexivity. + (* autobatch paramodulation by symm_andb,symm_orb,andb_true,Hletin,Hletin1. *) | simplify in H2 ⊢ %; intros; lapply (H2 i); clear H2; - autobatch + autobatch paramodulation; | simplify in H2 ⊢ %; intros; lapply (H2 i); clear H2; @@ -508,7 +521,7 @@ lemma look_for_axiom: simplify; apply (ex_intro ? ? a3); apply (ex_intro ? ? a4); - autobatch + autobatch depth=4 | right; intro; apply (bool_elim ? (same_atom a (FAtom n1)));