]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/demo/propositional_sequent_calculus.ma
please, commit files with debug=false otherwise the distributed matita prints a ton...
[helm.git] / helm / software / matita / library / demo / propositional_sequent_calculus.ma
index 4498676a3ef10f8a793464138f94e2638e130730..32e11054c5d19bdb699f77578b4fa60fd1c83154 100644 (file)
@@ -163,11 +163,10 @@ let rec and_of_list l ≝
   | cons F l' ⇒ FAnd F (and_of_list l')
   ].
 
-alias id "Nil" = "cic:/matita/list/list.ind#xpointer(1/1/1)".
 let rec or_of_list l ≝
  match l with
-  [ Nil ⇒ FFalse
-  | Cons F l' ⇒ FOr F (or_of_list l')
+  [ nil ⇒ FFalse
+  | cons F l' ⇒ FOr F (or_of_list l')
   ].
 
 definition formula_of_sequent ≝
@@ -752,18 +751,18 @@ lemma completeness_base:
   [ decompose;
     rewrite > H2; clear H2;
     rewrite > H4; clear H4;
-    apply (exchangeL ? a1 a2 (FAtom a));
-    apply (exchangeR ? a3 a4 (FAtom a));
-    apply axiom 
+    apply (ExchangeL ? a1 a2 (FAtom a));
+    apply (ExchangeR ? a3 a4 (FAtom a));
+    apply Axiom 
   | elim (sizel_0_no_axiom_is_tautology t t1 H H1 H2);
      [ decompose;
        rewrite > H3;
-       apply (exchangeL ? a a1 FFalse);
-       apply falseL
+       apply (ExchangeL ? a a1 FFalse);
+       apply FalseL
      | decompose;
        rewrite > H3;
-       apply (exchangeR ? a a1 FTrue);
-       apply trueR
+       apply (ExchangeR ? a a1 FTrue);
+       apply TrueR
      ]
   ]
 qed.