]> matita.cs.unibo.it Git - helm.git/commitdiff
test/a.ma => tests/ng_tactics.ma, with nassert here and there
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Thu, 16 Apr 2009 14:00:41 +0000 (14:00 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Thu, 16 Apr 2009 14:00:41 +0000 (14:00 +0000)
helm/software/matita/tests/a.ma [deleted file]
helm/software/matita/tests/ng_tactics.ma [new file with mode: 0644]

diff --git a/helm/software/matita/tests/a.ma b/helm/software/matita/tests/a.ma
deleted file mode 100644 (file)
index e6dac49..0000000
+++ /dev/null
@@ -1,57 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||M||                                                             *)
-(*      ||A||       A project by Andrea Asperti                           *)
-(*      ||T||                                                             *)
-(*      ||I||       Developers:                                           *)
-(*      ||T||         The HELM team.                                      *)
-(*      ||A||         http://helm.cs.unibo.it                             *)
-(*      \   /                                                             *)
-(*       \ /        This file is distributed under the terms of the       *)
-(*        v         GNU General Public License Version 2                  *)
-(*                                                                        *)
-(**************************************************************************)
-
-axiom A : Prop.
-axiom B : Prop.
-axiom C : Prop.
-axiom a: A.
-axiom b: B.
-
-definition xxx ≝ A.
-
-notation "#" non associative with precedence 90 for @{ 'sharp }.
-interpretation "a" 'sharp = a.
-interpretation "b" 'sharp = b.
-
-definition k : A → A ≝ λx.a.
-definition k1 : A → A ≝ λx.a.
-
-axiom P : A → Prop.
-
-include "nat/plus.ma".
-
-ntheorem pappo : ∀n:nat.n = S n + n → S n = (S (S n)).
-#m; #H; napply (let pippo ≝ (S m) in ?);
-nchange in match (S m) in ⊢ (?) with pippo;  
-
-nletin pippo ≝ (S m) in H ⊢ (?); 
-
-nwhd in H:(???%); 
-nchange in match (S ?) in H:% ⊢ (? → %) with (pred (S ?));  
-ntaint;
-
-ngeneralize in match m in ⊢ %;   in ⊢ (???% → ??%?);
-STOP 
-ncases (O) in m : % (*H : (??%?)*) ⊢ (???%);
-nelim (S m) in H : (??%?) ⊢ (???%);
-STOP;
-
-ntheorem pippo : ∀x:A. P (k x).
-nchange in match (k x) in ⊢ (∀_.%) with (k1 x); STOP
-
-ntheorem prova : (A → A → C) → C.
-napply (λH.?);
-napply (H ? ?); nchange A xxx; 
-napply #.
-nqed. 
\ No newline at end of file
diff --git a/helm/software/matita/tests/ng_tactics.ma b/helm/software/matita/tests/ng_tactics.ma
new file mode 100644 (file)
index 0000000..a9cb855
--- /dev/null
@@ -0,0 +1,105 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+include "nat/plus.ma".
+
+ntheorem test1 : ∀n:nat.n = S n + n → S n = (S (S n)).
+#m; #H;
+                                 nassert m:nat H: (m=S m + m) ⊢ (S m = S (S m));
+nletin pippo ≝ (S m) in H: % ⊢ (???%);
+ nassert m:nat pippo : nat ≝ (S m) ⊢ (m = pippo + m → (S m) = (S pippo));
+STOP;
+
+nwhd in H:(???%); 
+nchange in match (S ?) in H:% ⊢ (? → %) with (pred (S ?));  
+ntaint;
+
+ngeneralize in match m in ⊢ %;   in ⊢ (???% → ??%?);
+STOP 
+ncases (O) in m : % (*H : (??%?)*) ⊢ (???%);
+nelim (S m) in H : (??%?) ⊢ (???%);
+STOP;
+axiom A : Prop.
+axiom B : Prop.
+axiom C : Prop.
+axiom a: A.
+axiom b: B.
+
+include "nat/plus.ma".
+(*
+ntheorem foo: ∀n. n+n=n → n=n → n=n.
+ #n; #H; #K; nrewrite < H in (*K: (???%) ⊢*) ⊢ (??%?); napply (eq_ind ????? H);
+*)
+include "logic/connectives.ma".
+
+definition xxx ≝ A.
+
+notation "†" non associative with precedence 90 for @{ 'sharp }.
+interpretation "a" 'sharp = a.
+interpretation "b" 'sharp = b.
+
+include "nat/plus.ma".
+
+(*ntheorem foo: ∀n:nat. match n with [ O ⇒ n | S m ⇒ m + m ] = n.*)
+
+(*ntheorem prova : ((A ∧ A → A) → (A → B) → C) → C.
+# H; nassert H: ((A ∧ A → A) → (A → B) → C) ⊢ C;
+napply (H ? ?); nassert H: ((A ∧ A → A) → (A → B) → C) ⊢ (A → B)
+                        H: ((A ∧ A → A) → (A → B) → C) ⊢ (A ∧ A → A);
+ ##[#L | *; #K1; #K2]
+definition k : A → A ≝ λx.a.
+definition k1 : A → A ≝ λx.a.
+*)
+axiom P : A → Prop.
+
+include "nat/plus.ma".
+
+ntheorem pappo : ∀n:nat.n = S n + n → S n = (S (S n)).
+#m; #H; napply (let pippo ≝ (S m) in ?);
+nchange in match (S m) in ⊢ (?) with pippo;
+STOP (BUG)
+
+nletin pippo ≝ (S m) in H ⊢ (?); 
+
+nwhd in H:(???%); 
+nchange in match (S ?) in H:% ⊢ (? → %) with (pred (S ?));  
+ntaint;
+
+ngeneralize in match m in ⊢ %;   in ⊢ (???% → ??%?);
+STOP 
+ncases (O) in m : % (*H : (??%?)*) ⊢ (???%);
+nelim (S m) in H : (??%?) ⊢ (???%);
+STOP;
+
+ntheorem pippo : ∀x:A. P (k x).
+nchange in match (k x) in ⊢ (∀_.%) with (k1 x); STOP
+
+ntheorem prova : (A → A → C) → C.
+napply (λH.?);
+napply (H ? ?); nchange A xxx; 
+napply †.
+nqed. 
+
+REFL
+
+G
+
+{ r1 : T1, ..., r2 : T2 }
+
+reflexivity  apply REFL
+  =
+  apply (eq_refl ?);
+  apply (...)
+  apply (reflexivite S)
+  ...