X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Ftopology%2Figft2.ma;h=dc78b768079ef614c53908f3b3639d815c5fbcfd;hb=e880d6eab5e1700f4a625ddcd7d0fa8f0cce2dcc;hp=4ca6460d298aa7c25c3aefce96b6a1a86a21dfd4;hpb=31d160587c8d15e87959999025713395840c0b26;p=helm.git diff --git a/helm/software/matita/nlibrary/topology/igft2.ma b/helm/software/matita/nlibrary/topology/igft2.ma index 4ca6460d2..dc78b7680 100644 --- a/helm/software/matita/nlibrary/topology/igft2.ma +++ b/helm/software/matita/nlibrary/topology/igft2.ma @@ -23,7 +23,7 @@ ndefinition natS ≝ S. include "topology/igft.ma". nlemma hint_auto2 : ∀T.∀U,V:Ω^T.(∀x.x ∈ U → x ∈ V) → U ⊆ V. -nnormalize; nauto; +nnormalize; /2/; nqed. alias symbol "covers" = "covers set". @@ -33,9 +33,7 @@ nlemma cover_ind': ∀A:Ax.∀U,P:Ω^A. (U ⊆ P) → (∀a:A.∀j:𝐈 a. 𝐂 a j ◃ U → 𝐂 a j ⊆ P → a ∈ P) → ◃ U ⊆ P. - #A; #U; #P; #refl; #infty; #a; #H; nelim H - [ nauto | (*nauto depth=4;*) #b; #j; #K1; #K2; - napply infty; nauto; ##] + #A; #U; #P; #refl; #infty; #a; #H; nelim H; /3/. nqed. alias symbol "covers" (instance 1) = "covers". @@ -78,7 +76,7 @@ ncoercion uuax : ∀u:uuAx. Ax ≝ uuax on _u : uuAx to Ax. nlemma eq_rect_Type0_r': ∀A.∀a,x.∀p:eq ? x a.∀P: ∀x:A. eq ? x a → Type[0]. P a (refl A a) → P x p. - #A; #a; #x; #p; ncases p; nauto; + #A; #a; #x; #p; ncases p; //; nqed. nlemma eq_rect_Type0_r: @@ -108,24 +106,24 @@ nlet rec cover_rect ≝ ?. nlapply (decide_mem_ok … memdec b); nlapply (decide_mem_ko … memdec b); ncases (decide_mem … memdec b) - [ #_; #H; napply refl; nauto - | #H; #_; ncut (uuC … b=uuC … b) [nauto] ncases (uuC … b) in ⊢ (???% → ?) - [ #E; napply False_rect_Type0; ncut (b=b) [nauto] ncases p in ⊢ (???% → ?) - [ #a; #K; #E2; napply H [ nauto | nrewrite > E2; nauto ] + [ #_; #H; napply refl; /2/ + | #H; #_; ncut (uuC … b=uuC … b) [//] ncases (uuC … b) in ⊢ (???% → ?) + [ #E; napply False_rect_Type0; ncut (b=b) [//] ncases p in ⊢ (???% → ?) + [ #a; #K; #E2; napply H [ // | nrewrite > E2; // ] ##| #a; #i; #K; #E2; nrewrite < E2 in i; nnormalize; nrewrite > E; nnormalize; - nauto] + //] ##| #a; #E; ncut (a ◃ U) - [ nlapply E; nlapply (H ?) [nauto] ncases p + [ nlapply E; nlapply (H ?) [//] ncases p [ #x; #Hx; #K1; #_; ncases (K1 Hx) ##| #x; #i; #Hx; #K1; #E2; napply Hx; ngeneralize in match i; nnormalize; - nrewrite > E2; nnormalize; #_; nauto]##] + nrewrite > E2; nnormalize; #_; //]##] #Hcut; nlapply (infty b); nnormalize; nrewrite > E; nnormalize; #H2; napply (H2 one); #y; #E2; nrewrite > E2 - (* [##2: napply cover_rect] nauto depth=1; *) + (* [##2: napply cover_rect] //; *) [ napply Hcut - ##| napply (cover_rect A U memdec P refl infty a); nauto ]##] + ##| napply (cover_rect A U memdec P refl infty a); // ]##] nqed. (********* Esempio: @@ -171,6 +169,5 @@ ndefinition skipfact: ∀n:nat. n ◃ mk_powerclass ? (λx: uuax skipfact_dom.x= napply (S m * H (pred m) …); //] nqed. -nlemma test: skipfact four ? = eight. ##[##2: napply (skipfact_partial two)] - nnormalize; //. +nlemma test: skipfact four ? = eight. ##[##2: napply (skipfact_partial two)] //. nqed. \ No newline at end of file