X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Ftests%2Fcoercions_propagation.ma;h=7e3536b07071fc38d9192c0d138e8df71454b7ad;hb=e085135177f7b3b74b410d47a4f3bca1784b60b1;hp=c8529631b7a57515a4a2c59714ed634e30050982;hpb=751e50a8640e460a3d1a9fad96b8ec5fb3e663d0;p=helm.git diff --git a/helm/software/matita/tests/coercions_propagation.ma b/helm/software/matita/tests/coercions_propagation.ma index c8529631b..7e3536b07 100644 --- a/helm/software/matita/tests/coercions_propagation.ma +++ b/helm/software/matita/tests/coercions_propagation.ma @@ -12,28 +12,26 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/test/coercions_propagation/". + include "logic/connectives.ma". include "nat/orders.ma". -alias num (instance 0) = "natural number". inductive sigma (A:Type) (P:A → Prop) : Type ≝ sigma_intro: ∀a:A. P a → sigma A P. interpretation "sigma" 'exists \eta.x = - (cic:/matita/test/coercions_propagation/sigma.ind#xpointer(1/1) _ x). + (cic:/matita/tests/coercions_propagation/sigma.ind#xpointer(1/1) _ x). definition inject ≝ λP.λa:nat.λp:P a. sigma_intro ? P ? p. -coercion cic:/matita/test/coercions_propagation/inject.con 0 1. +coercion cic:/matita/tests/coercions_propagation/inject.con 0 1. definition eject ≝ λP.λc: ∃n:nat.P n. match c with [ sigma_intro w _ ⇒ w]. -coercion cic:/matita/test/coercions_propagation/eject.con. +coercion cic:/matita/tests/coercions_propagation/eject.con. alias num (instance 0) = "natural number". - theorem test: ∃n. 0 ≤ n. apply (S O : ∃n. 0 ≤ n). autobatch. @@ -50,16 +48,91 @@ qed. theorem test4: (∃n. 1 ≤ n) → ∃n. 0 < n. apply ((λn:nat.n) : (∃n. 1 ≤ n) → ∃n. 0 < n); - cases name_con; + cases s; assumption. qed. -(* -theorem test5: nat → ∃n. 0 ≤ n. - apply (λn:nat.?); - apply - (match n return λ_.∃n.0 ≤ n with [O ⇒ (0 : ∃n.0 ≤ n) | S n' ⇒ ex_intro ? ? n' ?] - : ∃n. 0 ≤ n); - autobatch. +theorem test5: nat → ∃n. 1 ≤ n. +apply ( + let rec aux n : nat ≝ + match n with + [ O ⇒ 1 + | S m ⇒ S (aux m) + ] + in + aux +: nat → ∃n. 1 ≤ n); +[ cases (aux n1); simplify; ] autobatch; +qed. + +inductive NN (A:Type) : nat -> Type ≝ + | NO : NN A O + | NS : ∀n:nat. NN A n → NN A (S n). + +definition injectN ≝ λA,k.λP.λa:NN A k.λp:P a. sigma_intro ? P ? p. + +coercion cic:/matita/tests/coercions_propagation/injectN.con 0 1. + +definition ejectN ≝ λA,k.λP.λc: ∃n:NN A k.P n. match c with [ sigma_intro w _ ⇒ w]. + +coercion cic:/matita/tests/coercions_propagation/ejectN.con. + +definition PN := + λA,k.λx:NN A k. 1 <= k. + +theorem test51_byhand: ∀A,k. NN A k → ∃n:NN A (S k). PN ? ? n. +intros 1; +apply ( + let rec aux (w : nat) (n : NN A w) on n : ∃p:NN A (S w).PN ? ? p ≝ + match n in NN return λx.λm:NN A x.∃p:NN A (S x).PN ? ? p with + [ NO ⇒ injectN ? ? ? (NS A ? (NO A)) ? + | NS x m ⇒ injectN ? ? ? (NS A (S x) (ejectN ? ? ? (aux ? m))) ? + ] + in + aux +: ∀n:nat. NN A n → ∃m:NN A (S n). PN ? ? m); +[2: cases (aux x m); simplify; autobatch ] unfold PN; autobatch; qed. -*) + +theorem f : nat -> nat -> ∃n:nat.O <= n. +apply (λx,y:nat.y : nat -> nat -> ∃n:nat. O <= n). +apply le_O_n; +qed. + +axiom F : nat -> nat -> nat. + +theorem f1 : nat -> nat -> ∃n:nat.O <= n. +apply (F : nat -> nat -> ∃n:nat. O <= n). +apply le_O_n; +qed. + +theorem test51: ∀A,k. NN A k → ∃n:NN A (S k). PN ? ? n. +intros 1; +letin xxx ≝ ( + let rec aux (w : nat) (n : NN A w) on n : NN A (S w) ≝ + match n in NN return λx.λm:NN A x.NN A (S x) with + [ NO ⇒ NS A ? (NO A) + | NS x m ⇒ NS A (S x) (aux ? m) + ] + in + aux +: ∀n:nat. NN A n → ∃m:NN A (S n). PN ? ? m); [3: apply xxx]; +unfold PN in aux ⊢ %; [cases (aux n2 n3)] autobatch; +qed. + +(* guarded troppo debole *) +theorem test522: nat → ∃n. 1 ≤ n. +apply ( + let rec aux n : nat ≝ + match n with + [ O ⇒ 1 + | S m ⇒ S (aux m) + ] + in + aux +: nat → ∃n. 1 ≤ n); +[ cases (aux n1); simplify; + autobatch +| autobatch]. +qed. +