]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/demo/propositional_sequent_calculus.ma
nasty change in the lexer/parser:
[helm.git] / helm / software / matita / library / demo / propositional_sequent_calculus.ma
index fb5278d6bdd78d44646652e6111317dc0fe3e370..42ae9682532ce8f65811f23944ec7e0f2a83cf0c 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/demo/propositional_sequent_calculus/".
-
 include "nat/plus.ma".
 include "nat/compare.ma".
 include "list/sort.ma".
@@ -183,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;
@@ -244,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;
@@ -258,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;
@@ -510,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)));