]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/sets/partitions.ma
Injectivity proved! What a mess...
[helm.git] / helm / software / matita / nlibrary / sets / partitions.ma
index f62c7281d1b074a068369cb9d71adfcca1c7b76d..acdd6a687cf49874f18f2fd23206bec5043d398b 100644 (file)
@@ -25,6 +25,9 @@ alias symbol "eq" = "setoid eq".
 alias symbol "eq" = "setoid1 eq".
 alias symbol "eq" = "setoid eq".
 alias symbol "eq" = "setoid1 eq".
+alias symbol "eq" = "setoid eq".
+alias symbol "eq" = "setoid1 eq".
+alias symbol "eq" = "setoid eq".
 nrecord partition (A: setoid) : Type[1] ≝ 
  { support: setoid;
    indexes: qpowerclass support;
@@ -65,7 +68,30 @@ naxiom minus_lt_to_lt: ∀n,m,p. n < p → minus n m < p.
 naxiom minus_O_n: ∀n. O = minus O n.
 naxiom le_O_to_eq: ∀n. n ≤ O → n=O.
 naxiom lt_to_minus_to_S: ∀n,m. m < n → ∃k. minus n m = S k.
-naxiom ltb_t: ∀n,m. n < m → ltb n m = true. 
+naxiom ltb_t: ∀n,m. n < m → ltb n m = true.
+naxiom ltb_f: ∀n,m. ¬ (n < m) → ltb n m = false.
+naxiom plus_n_O: ∀n. plus n O = n.
+naxiom not_lt_plus: ∀n,m. ¬ (plus n m < n).
+naxiom lt_to_lt_plus: ∀n,m,l. n < m → n < m + l.
+naxiom S_plus: ∀n,m. S (n + m) = n + S m.
+naxiom big_plus_ext: ∀n,f,f'. (∀i,p. f i p = f' i p) → big_plus n f = big_plus n f'.
+naxiom ad_hoc1: ∀n,m,l. n + (m + l) = l + (n + m).
+naxiom assoc: ∀n,m,l. n + m + l = n + (m + l).
+naxiom lt_canc: ∀n,m,p. n < m → p + n < p + m.
+naxiom ad_hoc2: ∀a,b. a < b → b - a - (b - S a) = S O.
+naxiom ad_hoc3: ∀a,b. b < a → S (O + (a - S b) + b) = a.
+naxiom ad_hoc4: ∀a,b. a - S b ≤ a - b.
+naxiom ad_hoc5: ∀a. S a - a = S O.
+naxiom ad_hoc6: ∀a,b. b ≤ a → a - b + b = a.
+naxiom ad_hoc7: ∀a,b,c. a + (b + O) + c - b = a + c.
+naxiom ad_hoc8: ∀a,b,c. ¬ (a + (b + O) + c < b).
+               
+
+naxiom split_big_plus:
+  ∀n,m,f. m ≤ n →
+   big_plus n f = big_plus m (λi,p.f i ?) + big_plus (n - m) (λi.λp.f (i + m) ?).
+ nelim daemon.
+nqed.
 
 nlemma partition_splits_card:
  ∀A. ∀P:partition A. ∀n,s.
@@ -90,20 +116,50 @@ nlemma partition_splits_card:
     ngeneralize in match (f_sur ???? (fi nindex) y ?) in ⊢ ?
      [##2: napply (. #‡(†?));##[##3: napply Hni2 |##2: ##skip | nassumption]##]
     *; #nindex2; *; #Hni21; #Hni22;
-    nletin xxx ≝ (plus match minus n nindex return λ_.nat with [ O ⇒ O | S nn ⇒ big_plus nn (λi.λ_.s (S (plus i nindex)))] nindex2);
+    nletin xxx ≝ (plus (big_plus (minus n nindex) (λi.λ_.s (S (plus i nindex)))) nindex2);
     napply (ex_intro … xxx); napply conj
-     [ nwhd in Hni1; nwhd; nelim daemon
-     | nwhd in ⊢ (???%?);
+     [ nwhd in Hni1; nwhd; nwhd in ⊢ (?(? %)%);
+       nchange with (? < plus (s n) (big_plus n ?));
+       nelim (le_to_lt_or_eq … (le_S_S_to_le … Hni1))
+        [##2: #E; nrewrite < E; nrewrite < (minus_canc nindex);
+          nwhd in ⊢ (?%?); nrewrite < E; napply lt_to_lt_plus; nassumption
+        | #L; nrewrite > (split_big_plus n (S nindex) (λm.λ_.s m) L);
+          nrewrite > (split_big_plus (n - nindex) (n - S nindex) (λi.λ_.s (S (i+nindex))) ?)
+           [ ngeneralize in match (big_plus_ext (n - S nindex)
+              (λi,p.s (S (i+nindex))) (λi,p.s (i + S nindex)) ?) in ⊢ ?
+               [ #E;
+                 napply (eq_rect_CProp0_r ??
+                  (λx:nat.λ_. x + big_plus (n - nindex - (n - S nindex))
+                   (λi,p.s (S (i + (n - S nindex)+nindex))) + nindex2 <
+                   s n + (big_plus (S nindex) (λi,p.s i) +
+                    big_plus (n - S nindex) (λi,p. s (i + S nindex)))) ? ? E);
+                 nrewrite > (ad_hoc1 (s n) (big_plus (S nindex) (λi,p.s i))
+                  (big_plus (n - S nindex) (λi,p. s (i + S nindex))));
+                 napply (eq_rect_CProp0_r
+                  ?? (λx.λ_.x < ?) ?? (assoc
+                  (big_plus (n - S nindex) (λi,p.s (i + S nindex)))
+                  (big_plus (n - nindex - (n - S nindex))
+                   (λi,p.s (S (i + (n - S nindex)+nindex))))
+                  nindex2));
+                 napply lt_canc;
+                 nrewrite > (ad_hoc2 … L); nwhd in ⊢ (?(?%?)?);
+                 nrewrite > (ad_hoc3 … L);
+                 napply (eq_rect_CProp0_r ?? (λx.λ_.x < ?) ?? (assoc …));
+                 napply lt_canc; nnormalize in ⊢ (?%?); nwhd in ⊢ (??%);
+                 napply lt_to_lt_plus; nassumption
+             ##|##2: #i; #_; nrewrite > (S_plus i nindex); napply refl]
+         ##| napply ad_hoc4]##]
+   ##| nwhd in ⊢ (???%?);
        nchange in Hni1 with (nindex < S n);
        ngeneralize in match (le_S_S_to_le … Hni1) in ⊢ ?;
        nwhd in ⊢ (? → ???(???????%?)?);
        napply (nat_rect_CProp0
         (λx. nindex ≤ x →
-          partition_splits_card_map A P (S n) s f fi
+          eq_rel (carr A) (eq A)
+          (partition_splits_card_map A P (S n) s f fi
            (plus
-             match minus x nindex with
-              [ O ⇒ O | S nn ⇒ big_plus nn (λi.λ_.s (S (plus i nindex)))]
-             nindex2) x = y) ?? n)
+             (big_plus (minus x nindex) (λi.λ_:i < minus x nindex.s (S (plus i nindex))))
+             nindex2) x) y) ?? n)
         [ #K; nrewrite < (minus_O_n nindex); nwhd in ⊢ (???(???????%?)?);
           nwhd in ⊢ (???%?); nchange in Hni21 with (nindex2 < s nindex);
           ngeneralize in match (le_O_to_eq … K) in ⊢ ?; #K';
@@ -115,119 +171,34 @@ nlemma partition_splits_card:
         | #n'; #Hrec; #HH; nelim (le_to_lt_or_eq … HH)
            [##2: #K; nrewrite < K; nrewrite < (minus_canc nindex);
             nwhd in ⊢ (???(???????%?)?);
-            (*???????*)
-         ##| #K; nwhd in ⊢ (???%?);
-             nrewrite > (minus_S n' nindex ?) [##2: napply le_S_S_to_le; nassumption]
-             ngeneralize in match (? :
-              match S (minus n' nindex) with [O ⇒ O | S nn ⇒ big_plus nn (λi.λ_.s (S (plus i nindex)))]
-              = big_plus (minus n' nindex) (λi.λ_.s (S (plus i nindex)))) in ⊢ ? [##2: napply refl]
-             #He; napply (eq_rect_CProp0_r ??
-              (λx.λ_.
-                match ltb (plus x nindex2) (s (S n')) with
-                 [ true ⇒ iso_f ???? (fi (S n')) (plus x nindex2)
-                 | false ⇒ ?(*partition_splits_card_map A P (S n) s f fi
-                     (minus (plus x nindex2) (s (S n'))) n'*)
-                 ] = y)
-              ?? He);
-             ngeneralize in match (? :
-              ltb (plus (big_plus (minus n' nindex) (λi.λ_.s (S (plus i nindex)))) nindex2)
-               (s (S n')) = false) in ⊢ ?
-               [ #Hc; nrewrite > Hc; nwhd in ⊢ (???%?);
-                 nelim (le_to_lt_or_eq … (le_S_S_to_le … K))
-                  [
-                ##| #E; ngeneralize in match Hc in ⊢ ?;
-                    nrewrite < E; nrewrite < (minus_canc nindex);
-                    nwhd in ⊢ (??(?%?)? → ?);
-                    nrewrite > E in Hni21; #E'; nchange in E' with (nindex2 < s n'); 
-                    ngeneralize in match Hni21 in ⊢ ?;
-                    
-               
-                ngeneralize in match (? :
-                 minus (plus (big_plus (minus n' nindex) (λi.λ_.s (S (plus i nindex)))) nindex2)
-                  (s (S n'))
-                 =
-                  plus
-                   match minus n' nindex with
-                    [ O ⇒ O | S nn ⇒ big_plus nn (λi.λ_.s (S (plus i nindex)))] nindex2)
-                 in ⊢ ?
-                  [ #F; nrewrite > F; napply Hrec; napply le_S_S_to_le; nassumption
-                  | nelim (le_to_lt_or_eq … (le_S_S_to_le … K))
-                     [
-                   ##| #E; nrewrite < E; nrewrite < (minus_canc nindex); nnormalize;
-                       
-                       nwhd in ⊢ (???%);
-                     ]
-             
-              nrewrite > He;
-             
-               
-              nnormalize in ⊢ (???%?);
-          
-       
-       
-       nelim (le_to_lt_or_eq … K)
-        [##2: #K'; nrewrite > K'; nrewrite < (minus_canc n); nnormalize;
-         napply (eq_rect_CProp0 nat nindex (λx:nat.λ_.partition_splits_card_map A P (S n) s f fi nindex2 x = y) ? n K');
-         nchange in Hni21 with (nindex2 < s nindex); ngeneralize in match Hni21 in ⊢ ?;
-         ngeneralize in match Hni22 in ⊢ ?;
-         nelim nindex
-          [ #X1; #X2; nwhd in ⊢ (??? % ?);
-            napply (lt_to_ltb_t ???? X2); #D; nwhd in ⊢ (??? % ?); nassumption
-          | #n0; #_; #X1; #X2; nwhd in ⊢ (??? % ?);
-            napply (lt_to_ltb_t ???? X2); #D; nwhd in ⊢ (??? % ?); nassumption]
-      ##| #K'; ngeneralize in match (lt_to_minus … K') in ⊢ ?; #K2;
-          napply (eq_rect_CProp0 ?? (λx.λ_.?) ? ? K2); (* uffa, ancora??? *)
-          nwhd in ⊢ (??? (???????(?%?)?) ?);
-          ngeneralize in match K' in ⊢ ?;
-          napply (nat_rect_CProp0
-           (λx. nindex < x →
-             partition_splits_card_map A P (S n) s f fi
-              (plus (big_op plus_magma_type (minus (minus x nindex) (S O))
-                (λi.λ_.s (S (plus i nindex))) O) nindex2) x = y) ?? n)
-           [ #A; nelim (not_lt_O … A)
-           | #n'; #Hrec; #X; nwhd in ⊢ (???%?);
-             ngeneralize in match
-              (? : ¬ ((plus (big_op plus_magma_type (minus (minus (S n') nindex) (S O))
-                (λi.λ_.s (S (plus i nindex))) O) nindex2) < s (S n'))) in ⊢ ?
-              [ #B1; napply (lt_to_ltb_f ???? B1); #B1'; nwhd in ⊢ (???%?);
-                nrewrite > (minus_S n' nindex …) [##2: napply le_S_S_to_le; nassumption]
-                ngeneralize in match (le_S_S_to_le … X) in ⊢ ?; #X';
-                nelim (le_to_lt_or_eq … X')
-                 [##2: #X'';
-                  nchange in Hni21 with (nindex2 < s nindex); ngeneralize in match Hni21 in ⊢ ?;
-                  nrewrite > X''; nrewrite < (minus_canc n');
-                  nrewrite < (minus_canc (S O)); nnormalize in ⊢ (? → %);
-                  nelim n'
-                   [ #Y; nwhd in ⊢ (??? % ?);
-                     ngeneralize in match (minus_lt_to_lt ? (s (S O)) ? Y) in ⊢ ?; #Y';
-                     napply (lt_to_ltb_t … Y'); #H; nwhd in ⊢ (???%?);  
-                 
-                nrewrite > (minus_S (minus n' nindex) (S O) …) [##2: 
-              
-              XXX;
-          
-          nelim n in f K' ⊢ ?
-           [ #A; nelim daemon;
-         
-       (* BEL POSTO DOVE FARE UN LEMMA *)
-       (* invariante: Hni1; altre premesse: Hni1, Hni22 *)
-       nelim n in ⊢ (% → ??? (????????%) ?)
-        [ #A (* decompose *)
-        | #index'; #Hrec; #K; nwhd in ⊢ (???%?);
-          nelim (ltb xxx (s (S index')));
-          #K1; nwhd in ⊢ (???%?)
-           [
-           
-           nindex < S index' + 1
-           +^{nindex} (s i) w < s (S index') 
-           S index' == nindex
-           
-           |
-           ]
-        ]
-     ]
-  | #x; #x'; nnormalize in ⊢ (? → ? → %);
+            nrewrite > K;
+            nwhd in ⊢ (???%?); nrewrite < K; nrewrite > (ltb_t … Hni21);
+            nwhd in ⊢ (???%?); nassumption
+         ##| #K; ngeneralize in match (le_S_S_to_le … K) in ⊢ ?; #K';
+             nwhd in ⊢ (???%?);
+             ngeneralize in match (?:
+              ¬ (big_plus (S n' - nindex) (λi,p.s (S (i+nindex))) + nindex2 < s (S n'))) in ⊢ ?
+              [ #N; nrewrite > (ltb_f … N); nwhd in ⊢ (???%?);
+                ngeneralize in match (Hrec K') in ⊢ ?; #Hrec';
+                napply (eq_rect_CProp0_r ??
+                 (λx,p. eq_rel (carr A) (eq A) (partition_splits_card_map A P (S n) s f fi
+                  (big_plus x ? + ? - ?) n') y) ?? (minus_S n' nindex K'));
+                nrewrite > (split_big_plus (S (n' - nindex)) (n' - nindex)
+                 (λi,p.s (S (i+nindex))) (le_S ?? (le_n ?)));
+                nrewrite > (ad_hoc5 (n' - nindex));
+                nnormalize in ⊢ (???(???????(?(?(??%)?)?)?)?);
+                nrewrite > (ad_hoc6 … K');
+                nrewrite > (ad_hoc7 (big_plus (n' - nindex) (λi,p.s (S (i+nindex))))
+                 (s (S n')) nindex2);
+                nassumption
+              | nrewrite > (minus_S … K');
+                nrewrite > (split_big_plus (S (n' - nindex)) (n' - nindex)
+                 (λi,p.s (S (i+nindex))) (le_S ?? (le_n ?)));
+                nrewrite > (ad_hoc5 (n' - nindex));
+                nnormalize in ⊢ (?(?(?(??%)?)?));
+                nrewrite > (ad_hoc6 … K');
+                napply ad_hoc8]##]##]##]
+##| #x; #x'; nnormalize in ⊢ (? → ? → %);
     nelim daemon
   ]
 nqed.
@@ -239,9 +210,11 @@ ndefinition partition_of_compatible_equivalence_relation:
  #A; #R; napply mk_partition
   [ napply (quotient ? R)
   | napply Full_set
-  | #a; napply mk_qpowerclass
-     [ napply {x | R x a}
-     | #x; #x'; #H; nnormalize; napply mk_iff; #K; nelim daemon]
+  | napply mk_unary_morphism1
+     [ #a; napply mk_qpowerclass
+        [ napply {x | R x a}
+        | #x; #x'; #H; nnormalize; napply mk_iff; #K; nelim daemon]
+   ##| #a; #a'; #H; napply conj; #x; nnormalize; #K [ nelim daemon | nelim daemon]##]
 ##| #x; #_; nnormalize; napply (ex_intro … x); napply conj; napply refl
   | #x; #x'; #_; #_; nnormalize; *; #x''; *; #H1; #H2; napply (trans ?????? H2);
     napply sym; nassumption