| 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)).
+
nlemma partition_splits_card:
∀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; #n; #s; #f; #fi; napply mk_isomorphism
+ #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 n s f fi m n)
+ [ 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 (f_sur ???? (fi nindex) y ?) in ⊢ ?
[##2: napply (. #‡(†?));##[##3: napply Hni2 |##2: ##skip | nassumption]##]
*; #nindex2; *; #Hni21; #Hni22;
- napply (ex_intro … (plus (big_plus nindex (λi.λ_.s i)) nindex2)); napply conj
+ nletin xxx ≝ (plus match nindex return λ_.nat with [ O ⇒ O | S nn ⇒ big_plus nn (λi.λ_.s i)] nindex2);
+ napply (ex_intro … xxx); napply conj
[ nwhd in Hni1; nelim daemon
| nwhd in ⊢ (???%?);
(* BEL POSTO DOVE FARE UN LEMMA *)
(* invariante: Hni1; altre premesse: Hni1, Hni22 *)
- nchange in Hni1 with (nindex < n); ngeneralize in match Hni1 in ⊢ ?;
- nelim n
+ nchange in Hni1 with (nindex < S n); ngeneralize in match Hni1 in ⊢ ?;
+ nelim n in ⊢ (% → ??? (????????%) ?)
[ #A (* decompose *)
- | #index'; #Hrec;
+ | #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 ⊢ (? → ? → %);