X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Fsets%2Fpartitions.ma;h=f62c7281d1b074a068369cb9d71adfcca1c7b76d;hb=d70944c1513aa63e6494e58595fcc4214a2f6c68;hp=43940e1f9cc4b1d7831a841fe9abddff81b1bbf5;hpb=c6d3537eee27d05490a9555cc7326bc954b356c5;p=helm.git diff --git a/helm/software/matita/nlibrary/sets/partitions.ma b/helm/software/matita/nlibrary/sets/partitions.ma index 43940e1f9..f62c7281d 100644 --- a/helm/software/matita/nlibrary/sets/partitions.ma +++ b/helm/software/matita/nlibrary/sets/partitions.ma @@ -17,49 +17,234 @@ include "nat/plus.ma". include "nat/compare.ma". include "nat/minus.ma". -(* sbaglia a fare le proiezioni *) -nrecord finite_partition (A: Type[0]) : Type[1] ≝ - { card: nat; - class: ∀n. lt n card → Ω \sup A; - inhabited: ∀i,p. class i p ≬ class i p(*; - disjoint: ∀i,j,p,p'. class i p ≬ class j p' → i=j; - covers: big_union ?? class = full_set A*) - }. +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". +alias symbol "eq" = "setoid1 eq". +nrecord partition (A: setoid) : Type[1] ≝ + { support: setoid; + indexes: qpowerclass support; + class: unary_morphism1 (setoid1_of_setoid support) (qpowerclass_setoid A); + inhabited: ∀i. i ∈ indexes → class i ≬ class i; + disjoint: ∀i,j. i ∈ indexes → j ∈ indexes → class i ≬ class j → i=j; + covers: big_union support ? ? (λx.class x) = full_set A + }. napply indexes; nqed. -nrecord has_card (A: Type[0]) (S: Ω \sup A) (n: nat) : CProp[0] ≝ - { f: ∀m:nat. lt m n → A; - in_S: ∀m.∀p:lt m n. f ? p ∈ S (*; - f_inj: injective ?? f; - f_sur: surjective ?? f*) - }. - -(* -nlemma subset_of_finite: - ∀A. ∃n. has_card ? (full_subset A) n → ∀S. ∃m. has_card ? S m. -nqed. -*) +naxiom daemon: False. nlet rec partition_splits_card_map - A (P: finite_partition A) (s: ∀i. lt i (card ? P) → nat) - (H:∀i.∀p: lt i (card ? P). has_card ? (class ? P i p) (s i p)) - m (H1: lt m (big_plus ? s)) index (p: lt index (card ? P)) on index : A ≝ - match index return λx. lt x (card ? P) → ? with - [ O ⇒ λp'.f ??? (H O p') m ? - | S index' ⇒ λp'. - match ltb m (s index p) with - [ or_introl K ⇒ f ??? (H index p) m K - | or_intror _ ⇒ partition_splits_card_map A P s H (minus m (s index p)) ? index' ? ]] p. -##[##2: napply lt_minus; nassumption - |##3: napply lt_Sn_m; nassumption - | -nqed. + A (P:partition A) n s (f:isomorphism ?? (Nat_ n) (indexes ? P)) + (fi: ∀i. isomorphism ?? (Nat_ (s i)) (class ? P (iso_f ???? f i))) m index + on index : A ≝ + match ltb m (s index) with + [ true ⇒ iso_f ???? (fi index) m + | false ⇒ + match index with + [ O ⇒ (* dummy value: it could be an elim False: *) iso_f ???? (fi O) O + | S index' ⇒ + partition_splits_card_map A P n s f fi (minus m (s index)) index']]. + +naxiom big_union_preserves_iso: + ∀A,A',B,T,T',f. + ∀g: isomorphism A' A T' T. + big_union A B T f = big_union A' B T' (λx.f (iso_f ???? g x)). + +naxiom le_to_lt_or_eq: ∀n,m. n ≤ m → n < m ∨ n = m. +alias symbol "eq" = "leibnitz's equality". +naxiom minus_canc: ∀n. O = minus n n. +naxiom lt_to_ltb_t: ∀n,m. ∀P: bool → CProp[0]. P true → n < m → P (ltb n m). +naxiom lt_to_ltb_f: ∀n,m. ∀P: bool → CProp[0]. P false → ¬ (n < m) → P (ltb n m). +naxiom lt_to_minus: ∀n,m. n < m → S (minus (minus m n) (S O)) = minus m n. +naxiom not_lt_O: ∀n. ¬ (n < O). +naxiom minus_S: ∀n,m. m ≤ n → minus (S n) m = S (minus n m). +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. nlemma partition_splits_card: - ∀A. ∀P: finite_partition A. ∀s: ∀i. lt i (card ? P) → nat. - (∀i.∀p: lt i (card ? P). has_card ? (class ? P i p) (s i p)) → - has_card A (full_set A) (big_plus (card ? P) s). - #A; #P; #s; #H; napply mk_has_card - [ #m; #H1; napply partition_splits_card_map A P s H m H1 (pred (card ? P)) - | + ∀A. ∀P:partition A. ∀n,s. + ∀f:isomorphism ?? (Nat_ n) (indexes ? P). + (∀i. isomorphism ?? (Nat_ (s i)) (class ? P (iso_f ???? f i))) → + (isomorphism ?? (Nat_ (big_plus n (λi.λ_.s i))) (Full_set A)). + #A; #P; #Sn; ncases Sn + [ #s; #f; #fi; + ngeneralize in match (covers ? P) in ⊢ ?; *; #_; #H; + ngeneralize in match + (big_union_preserves_iso ??? (indexes A P) (Nat_ O) (λx.class ? P x) f) in ⊢ ?; + *; #K; #_; nwhd in K: (? → ? → %); + nelim daemon (* impossibile *) + | #n; #s; #f; #fi; napply mk_isomorphism + [ napply mk_unary_morphism + [ napply (λm.partition_splits_card_map A P (S n) s f fi m n) + | #a; #a'; #H; nrewrite < H; napply refl ] +##| #y; #_; + ngeneralize in match (covers ? P) in ⊢ ?; *; #_; #Hc; + ngeneralize in match (Hc y I) in ⊢ ?; *; #index; *; #Hi1; #Hi2; + ngeneralize in match (f_sur ???? f ? Hi1) in ⊢ ?; *; #nindex; *; #Hni1; #Hni2; + 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); + napply (ex_intro … xxx); napply conj + [ nwhd in Hni1; nwhd; nelim daemon + | 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 + (plus + match minus x nindex with + [ O ⇒ O | S nn ⇒ big_plus nn (λi.λ_.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'; + ngeneralize in match Hni21 in ⊢ ?; + ngeneralize in match Hni22 in ⊢ ?; + nrewrite > K' in ⊢ (% → % → ?); #K1; #K2; + nrewrite > (ltb_t … K2); + nwhd in ⊢ (???%?); nassumption + | #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 ⊢ (? → ? → %); + nelim daemon ] +nqed. + +(************** equivalence relations vs partitions **********************) + +ndefinition partition_of_compatible_equivalence_relation: + ∀A:setoid. compatible_equivalence_relation A → partition A. + #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] +##| #x; #_; nnormalize; napply (ex_intro … x); napply conj; napply refl + | #x; #x'; #_; #_; nnormalize; *; #x''; *; #H1; #H2; napply (trans ?????? H2); + napply sym; nassumption + | nnormalize; napply conj + [ #a; #_; napply I | #a; #_; napply (ex_intro … a); napply conj [ napply I | napply refl]##] nqed. \ No newline at end of file