From de21be5819bd35a2cb83b3d33b1c578d970a32c7 Mon Sep 17 00:00:00 2001 From: marangon Date: Thu, 15 Dec 2005 16:03:01 +0000 Subject: [PATCH] New tactic: inversion. - only first phase implemented (no cleaning) - code in inversion.ml to be improved/cleaned up --- helm/matita/tests/inversion.ma | 68 ++++++++++++++++++--------------- helm/matita/tests/inversion2.ma | 44 +++++++++++++++++++++ 2 files changed, 82 insertions(+), 30 deletions(-) create mode 100644 helm/matita/tests/inversion2.ma diff --git a/helm/matita/tests/inversion.ma b/helm/matita/tests/inversion.ma index f717cd1df..b39d839b0 100644 --- a/helm/matita/tests/inversion.ma +++ b/helm/matita/tests/inversion.ma @@ -12,42 +12,50 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/tests/inversion/". +set "baseuri" "cic:/matita/tests/inversion_sum/". include "coq.ma". -inductive nat : Set \def - O : nat - | S : nat \to nat. - -inductive le (n:nat) : nat \to Prop \def - leO : le n n - | leS : \forall m. le n m \to le n (S m). alias symbol "eq" (instance 0) = "Coq's leibnitz's equality". +alias id "O" = "cic:/Coq/Init/Datatypes/nat.ind#xpointer(1/1/1)". + +inductive sum (n:nat) : nat \to nat \to Set \def + k: \forall x,y. n = x + y \to sum n x y. + + -theorem test_inversion: \forall n. le n O \to n=O. + +theorem t: \forall x,y. \forall H: sum x y O. + match H with [ (k a b p) \Rightarrow a ] = x. intros. - (* inversion begins *) - cut (O=O); - [ 2: reflexivity; - | generalize in match Hcut. - apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H); - [ intro. rewrite < H1. clear Hcut. - | simplify. intros. discriminate H3. - ] - reflexivity. - ] + inversion H. + + (* + cut (y = y \to O = O \to match H with [ (k a b p) \Rightarrow a] = x). + apply Hcut; reflexivity. + apply + (sum_ind ? + (\lambda a,b,K. y=a \to O=b \to + match K with [ (k a b p) \Rightarrow a ] = x) + ? ? ? H). + goal 16.*) + simplify. intros. + generalize in match H1. + rewrite < H2; rewrite < H3.intro. + rewrite > H4.auto. qed. -(* Piu' semplice e non lascia l'ipotesi inutile Hcut *) -alias id "refl_equal" = "cic:/Coq/Init/Logic/eq.ind#xpointer(1/1/1)". -theorem test_inversion2: \forall n. le n O \to n=O. - intros. - (* inversion begins *) - generalize in match (refl_equal nat O). - apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H); - [ intro. rewrite < H1. - | simplify. intros. discriminate H3. - ] - reflexivity. +theorem t1: \forall x,y. sum x y O \to x = y. +intros. + +(* +cut y=y \to O=O \to x = y. +apply Hcut.reflexivity. reflexivity. +apply (sum_ind ? (\lambda a,b,K. y=a \to O=b \to x=a) ? ? ? s).*) + +(*apply (sum_ind ? (\lambda a,b,K. y = a \to O = b \to x = a) ? ? ? s).*) +inversion s. +intros.simplify. +intros. +rewrite > H. rewrite < H2. auto. qed. diff --git a/helm/matita/tests/inversion2.ma b/helm/matita/tests/inversion2.ma new file mode 100644 index 000000000..bb8d18980 --- /dev/null +++ b/helm/matita/tests/inversion2.ma @@ -0,0 +1,44 @@ +(**************************************************************************) +(* ___ *) +(* ||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 *) +(* *) +(**************************************************************************) + +set "baseuri" "cic:/matita/tests/inversion/". +include "coq.ma". + +inductive nat : Set \def + O : nat + | S : nat \to nat. + + +inductive le (n:nat) : nat \to Prop \def + leO : le n n + | leS : \forall m. le n m \to le n (S m). + + +inductive ledx : nat \to nat \to Prop \def + ledxO : \forall n. ledx n n + | ledxS : \forall m.\forall n. ledx n m \to ledx n (S m). + + +alias symbol "eq" (instance 0) = "Coq's leibnitz's equality". + +theorem test_inversion: \forall n. le n O \to n=O. + intros. + inversion H. + (* cut n=n \to O=O \to n=O. + apply Hcut; reflexivity. *) + (* elim H. BUG DI UNSHARING *) + (*apply (ledx_ind (\lambda x.\lambda y. n=x \to O=y \to x=y) ? ? ? ? H).*) + simplify. intros. reflexivity. + simplify. intros. discriminate H3. +qed. -- 2.39.2