X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Ftopology%2Figft3.ma;h=e242cb882d52e527f812d2cbf89457885377c8df;hb=31d160587c8d15e87959999025713395840c0b26;hp=32d0f8f8469723db8d0e0ec8a7f2feaa29f6f66d;hpb=8b33918a53a2c752e2256b328bc0e45f4827beac;p=helm.git diff --git a/helm/software/matita/nlibrary/topology/igft3.ma b/helm/software/matita/nlibrary/topology/igft3.ma index 32d0f8f84..e242cb882 100644 --- a/helm/software/matita/nlibrary/topology/igft3.ma +++ b/helm/software/matita/nlibrary/topology/igft3.ma @@ -11,18 +11,16 @@ (* v GNU General Public License Version 2 *) (* *) (**************************************************************************) -(* + include "arithmetics/nat.ma". -*) include "datatypes/bool.ma". -(* ndefinition two ≝ S (S O). ndefinition natone ≝ S O. ndefinition four ≝ two * two. ndefinition eight ≝ two * four. ndefinition natS ≝ S. -*) + include "topology/igft.ma". nlemma hint_auto2 : ∀T.∀U,V:Ω^T.(∀x.x ∈ U → x ∈ V) → U ⊆ V. @@ -134,8 +132,6 @@ nrecord memdec (A: Type[0]) (U:Ω^A) : Type[0] ≝ (*********** end from Cantor ********) -naxiom daemon: False. - nlemma csc_sym_eq: ∀A,x,y. eq A x y → eq A y x. #A; #x; #y; #H; ncases H; @1. nqed. @@ -192,34 +188,23 @@ nlet rec cover_rect nqed. (********* Esempio: - let rec skipfact n = + let rec skip n = match n with [ O ⇒ 1 - | S m ⇒ S m * skipfact (pred m) ] + | S m ⇒ + match m with + [ O ⇒ skipfact O + | S _ ⇒ S m * skipfact (pred m) * skipfact (pred m) ]] **) -ntheorem psym_plus: ∀n,m. n + m = m + n. - #n; nelim n - [ #m; nelim m; //; #n0; #H; - nchange with (natS n0 = natS (n0 + O)); - nrewrite < H; // - | #n0; #H; #m; nchange with (S (n0 + m) = m + S n0); - nrewrite > (H …); - nelim m; //; - #n1; #E; nrewrite > E; //] +ntheorem psym_plus: ∀n,m. n + m = m + n.//. nqed. -nlemma easy1: ∀n:nat. two * (S n) = two + two * n. - #n; nelim n;//; - #n0; #H; nnormalize; - nrewrite > (psym_plus ??); - nrewrite > H; nnormalize; - nrewrite > (psym_plus ??); - //. +nlemma easy1: ∀n:nat. two * (S n) = two + two * n.//. nqed. ndefinition skipfact_dom: uuAx. - @ nat; #n; ncases n [ napply None | #m; napply (Some … (pred m)) ] + @ nat; #n; ncases n [ napply None | #m; ncases m [ napply (Some … O) | #_; napply (Twice … (pred m) (pred m)) ] nqed. ntheorem skipfact_base_dec: @@ -232,11 +217,14 @@ ntheorem skipfact_partial: ∀n: uuax skipfact_dom. two * n ◃ mk_powerclass ? (λx: uuax skipfact_dom.x=O). #n; nelim n [ @1; nnormalize; @1 - | #m; #H; @2 - [ nnormalize; @1 - | nnormalize; #y; nchange in ⊢ (% → ?) with (y = pred (pred (two * (natone + m)))); - nrewrite > (easy1 …); nchange in ⊢ (% → ?) with (y = two * m); - #E; nrewrite > E; nassumption ]##] + | #m; nelim m; nnormalize + [ #H; @2; nnormalize + [ // + | #y; *; #a; #E; nrewrite > E; ncases a; nnormalize; // ] + ##| #p; #H1; #H2; @2; nnormalize + [ // + | #y; *; #a; #E; nrewrite > E; ncases a; nnormalize; + nrewrite < (plus_n_Sm …); // ]##] nqed. ndefinition skipfact: ∀n:nat. n ◃ mk_powerclass ? (λx: uuax skipfact_dom.x=O) → nat. @@ -244,10 +232,13 @@ ndefinition skipfact: ∀n:nat. n ◃ mk_powerclass ? (λx: uuax skipfact_dom.x= [ #a; #_; napply natone | #a; ncases a [ nnormalize; #i; nelim i - | #m; #i; nnormalize in i; #d; #H; - napply (S m * H (pred m) …); //] + | #m; ncases m + [ nnormalize; #_; #_; #H; napply H; @1 + | #p; #i; nnormalize in i; #K; + #H; nnormalize in H; + napply (S m * H true * H false) ] nqed. -nlemma test: skipfact four ? = eight. ##[##2: napply (skipfact_partial two)] +nlemma test: skipfact four ? = four * two * two. ##[##2: napply (skipfact_partial two)] nnormalize; //. nqed. \ No newline at end of file