X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Ftests%2Fcoercions_russell.ma;h=c386743e74f466edcf5bff9938d9d9142b6de1de;hb=275b124e484510dc49141f86b5174f8bd0be7d97;hp=dede79ca145e741115813aac4ebff05b48cedb13;hpb=48b5f69a87f0f4e067c89085ddd1ae503e53e068;p=helm.git diff --git a/helm/software/matita/tests/coercions_russell.ma b/helm/software/matita/tests/coercions_russell.ma index dede79ca1..c386743e7 100644 --- a/helm/software/matita/tests/coercions_russell.ma +++ b/helm/software/matita/tests/coercions_russell.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/test/russell/". + include "nat/orders.ma". include "list/list.ma". @@ -22,15 +22,15 @@ inductive sigma (A:Type) (P:A → Prop) : Type ≝ sig_intro: ∀a:A. P a → sigma A P. interpretation "sigma" 'exists \eta.x = - (cic:/matita/test/russell/sigma.ind#xpointer(1/1) _ x). + (cic:/matita/tests/coercions_russell/sigma.ind#xpointer(1/1) _ x). definition inject ≝ λP.λa:list nat.λp:P a. sig_intro ? P ? p. -coercion cic:/matita/test/russell/inject.con 0 1. +coercion cic:/matita/tests/coercions_russell/inject.con 0 1. definition eject ≝ λP.λc: ∃n:list nat.P n. match c with [ sig_intro w _ ⇒ w]. -coercion cic:/matita/test/russell/eject.con. +coercion cic:/matita/tests/coercions_russell/eject.con. alias symbol "exists" (instance 2) = "exists". lemma tl : ∀l:list nat. l ≠ [] → ∃l1.∃a.a :: l1 = l. @@ -44,18 +44,18 @@ qed. alias symbol "exists" (instance 3) = "exists". lemma tl2 : ∀l:∃l:list nat. l ≠ []. ∃l1.∃a.a :: l1 = l. -letin program ≝ - (λl:list nat. match l with [ nil ⇒ [] | cons x l1 ⇒ l1]); -letin program_spec ≝ - (program : ∀l:∃l:list nat. l ≠ []. ∃l1.∃a.a :: l1 = l); - [ autobatch; | generalize in match H; clear H; cases s; simplify; - intros; cases (H H1); ] -exact program_spec. +apply + (λl:list nat. match l with [ nil ⇒ [] | cons x l1 ⇒ l1] + : + ∀l:∃l:list nat. l ≠ []. ∃l1.∃a.a :: l1 = l); +[ autobatch; +| generalize in match H; clear H; cases s; simplify; + intros; cases (H H1);] qed. definition nat_return := λn:nat. Some ? n. -coercion cic:/matita/test/russell/nat_return.con. +coercion cic:/matita/tests/coercions_russell/nat_return.con. definition raise_exn := None nat. @@ -76,11 +76,11 @@ include "nat/compare.ma". definition inject_opt ≝ λP.λa:option nat.λp:P a. sig_intro ? P ? p. -coercion cic:/matita/test/russell/inject_opt.con 0 1. +coercion cic:/matita/tests/coercions_russell/inject_opt.con 0 1. definition eject_opt ≝ λP.λc: ∃n:option nat.P n. match c with [ sig_intro w _ ⇒ w]. -coercion cic:/matita/test/russell/eject_opt.con. +coercion cic:/matita/tests/coercions_russell/eject_opt.con. (* we may define mem as in the following lemma and get rid of it *) definition find_spec ≝ @@ -93,7 +93,7 @@ definition find_spec ≝ ∃l1,l2,l3.l = l1 @ [x] @ l2 @ [y] @ l3]. lemma mem_x_to_ex_l1_l2 : ∀l,x.mem ? eqb x l = true → ∃l1,l2.l = l1 @ [x] @ l2. -intros 2; elim l (H hd tl IH H); [destruct H] +intros 2; elim l (H hd tl IH H); [simplify in H; destruct H] generalize in match H; clear H; simplify; apply (eqb_elim x hd); simplify; intros; [1:clear IH; rewrite < H; apply (ex_intro ? ? []); @@ -101,23 +101,19 @@ simplify; apply (eqb_elim x hd); simplify; intros; simplify; autobatch; qed. -definition if : ∀A:Type.bool → A → A → A ≝ - λA,b,t,f. match b with [ true ⇒ t | false ⇒ f]. - -notation < "'If' \nbsp b \nbsp 'Then' \nbsp t \nbsp 'Else' \nbsp f" for @{ 'if $b $t $f }. -notation > "'If' b 'Then' t 'Else' f" for @{ 'if $b $t $f }. -interpretation "if" 'if a b c = (cic:/matita/test/russell/if.con _ a b c). +notation > "'If' b 'Then' t 'Else' f" +for @{ match $b with [ true ⇒ $t | _ ⇒ $f ] }. definition sigma_find_spec : ∀p,l. sigma ? (λres.find_spec l p res). (* define the find function *) -letin find ≝ (λp. +apply (λp. let rec aux l ≝ match l with [ nil ⇒ raise_exn | cons x l ⇒ If p x Then nat_return x Else aux l] - in aux); -(* manca una delta?! *) unfold if in find; -apply (find: ∀p.∀l.sigma ? (λres.find_spec l p res)); clear find; + in aux + : + ∀p.∀l.sigma ? (λres.find_spec l p res)); (* l = x::tl ∧ p x = false *) [1: cases (aux l1); clear aux; generalize in match H2; clear H2; cases a; clear a; simplify;