Library sset9

Theory of Sets EIII-5 Properties of integers

Copyright INRIA (2009- 2011) Apics Team (Jose Grimm).

Require Import ssreflect ssrbool eqtype ssrnat ssrfun.
Set Implicit Arguments.
Unset Strict Implicit.
Import Prenex Implicits.
Require Export sset8.


Module IntegerProps.

Hint Resolve zero_smallest: fprops.

EIII-5-1 Operations on integers and finite sets


Functions on nat

Lemma Bsum_increasing3: forall a b, inc a Bnat -> inc b Bnat->
  a <=c (a +c b).
Proof.
move=> a b ca cb; apply sum_increasing3; fprops.
Qed.

Lemma Bprod_increasing3: forall a b, inc a Bnat -> inc b Bnat->
   b <> 0c -> a <=c (a *c b).
Proof.
move=> a b ca cb; apply product_increasing3; fprops.
Qed.

Lemma Bnat_total_order1: forall a b, inc a Bnat -> inc b Bnat ->
  a = b \/ a <c b \/ b<c a.
Proof. move=> a b aB bB; apply cardinal_le_total_order1; fprops. Qed.

Lemma Bnat_total_order2: forall a b, inc a Bnat -> inc b Bnat ->
  a <=c b \/ b <c a.
Proof. move=> a b aB bB; apply cardinal_le_total_order2; fprops. Qed.

Lemma trivial_cardinal_sum3: forall f a, fgraph f ->
  inc a (domain f) ->
  cardinal_sum (restr f (singleton a)) = cardinal (V a f).
Proof.
move=> f a fgf cV.
have ssd: sub (singleton a) (domain f).
  by move=> t; rewrite singleton_rw => ->.
have Va: V a (restr f (singleton a)) = V a f by bw; fprops.
rewrite -Va ; apply trivial_cardinal_sum2; rewrite restr_domain1 //.
Qed.

Lemma trivial_cardinal_prod3: forall f a, fgraph f ->
  inc a (domain f) ->
  cardinal_prod (restr f (singleton a)) = cardinal (V a f).
Proof.
move=> f a fgf adf.
have ssd: sub (singleton a) (domain f).
  by move=> t; rewrite singleton_rw => ->.
have Va: V a (restr f (singleton a)) = V a f by bw ; fprops.
rewrite -Va;apply trivial_cardinal_prod2; [ fprops | rewrite restr_domain1 //]. Qed.

Lemma partition_tack_on: forall a b, ~ inc b a ->
  partition_fam (variantLc a (singleton b)) (tack_on a b).
Proof.
move=> a b bna;red; ee.
  have d1: disjoint a (singleton b) by apply disjoint_with_singleton.
  move=> i j; bw; move=> it jt.
  try_lvariant it; try_lvariant jt;auto; right.
  by apply disjoint_symmetric.
rewrite /unionb variantLc_domain -two_points_pr2 union_of_twosets_aux.
by rewrite variant_V_ca variant_V_cb.
Qed.

Lemma partition_complement: forall a b, inc b a ->
  partition_fam (variantLc (complement a (singleton b)) (singleton b)) a.
Proof.
move=> a b ba.
set(c:= (complement a (singleton b))).
have nbc: ~ (inc b c) by rewrite /c; srw; aw; intuition.
have acb: a = tack_on c b by rewrite /c; apply tack_on_complement.
by rewrite acb; apply partition_tack_on.
Qed.

Lemma induction_sum0: forall f a b, fgraph f ->
  sub (tack_on a b) (domain f) -> (~ inc b a) ->
  cardinal_sum (restr f (tack_on a b)) =
  cardinal_sum (restr f a) +c (V b f).
Proof.
move=> f a b fgf sta nba.
move: (partition_tack_on nba) => pfa.
set (g:= variantLc a (singleton b)).
have pfg: partition_fam g (domain (restr f (tack_on a b))).
  by rewrite restr_domain1.
have fgr: fgraph (restr f(tack_on a b)) by fprops.
have aux1: sub a (tack_on a b) by fprops.
have aux2: sub (singleton b) (tack_on a b) by move=> t; aw; fprops.
rewrite (cardinal_sum_assoc fgr pfg) /g; bw; rewrite card_plus_pr0; bw.
have bdf: inc b (domain f) by fprops.
rewrite double_restr; first rewrite double_restr;
  first rewrite (trivial_cardinal_sum3 fgf bdf) card_plus_pr2b ; fprops.
Qed.

Lemma induction_prod0: forall f a b, fgraph f ->
  sub (tack_on a b) (domain f) -> (~ inc b a) ->
  cardinal_prod (restr f (tack_on a b)) =
  (cardinal_prod (restr f a)) *c (V b f) .
Proof.
move=> f a b fgf tadf nba.
move: (partition_tack_on nba) => pfa.
set (g:= variantLc a (singleton b)).
have pfg: partition_fam g (domain (restr f (tack_on a b))).
  by rewrite restr_domain1.
have fgr: fgraph (restr f (tack_on a b)) by fprops.
have aux1: sub a (tack_on a b) by fprops.
have aux2: sub (singleton b) (tack_on a b) by move=> t; aw; fprops.
rewrite (cardinal_prod_assoc fgr pfg) /g; bw; rewrite card_mult_pr0; bw.
have bdf: inc b (domain f) by fprops.
rewrite double_restr; first rewrite double_restr;
  first rewrite (trivial_cardinal_prod3 fgf bdf) card_mult_pr2b; fprops.
Qed.

Lemma induction_sum1: forall f a b, fgraph f ->
  domain f = tack_on a b -> (~ inc b a) ->
  cardinal_sum f =
  cardinal_sum (restr f a) +c (V b f).
Proof.
move=> f a b fgf df nba; rewrite - induction_sum0 =>//.
  apply f_equal; rewrite -df; symmetry;apply restr_to_domain; ee.
rewrite df;fprops.
Qed.

Lemma induction_prod1: forall f a b, fgraph f ->
  domain f = tack_on a b -> (~ inc b a) ->
  cardinal_prod f =
  cardinal_prod (restr f a) *c (V b f).
Proof.
move=> f a b fgf df nba; rewrite - induction_prod0 =>//.
  apply f_equal; rewrite -df; symmetry;apply restr_to_domain; ee.
rewrite df;fprops.
Qed.

Lemma sum_increasing6: forall f j, fgraph f ->
  (forall x, inc x (domain f) -> is_cardinal (V x f)) ->
  inc j (domain f) -> (V j f) <=c (cardinal_sum f).
Proof.
move=> f j fgf alc jd.
have sjd: (sub (singleton j) (domain f)).
  by move=> t td; awi td; rewrite td.
move: (sum_increasing1 fgf alc sjd).
rewrite trivial_cardinal_sum3 //.
by rewrite (cardinal_of_cardinal (alc _ jd)).
Qed.

Lemma prod_increasing6: forall f j, fgraph f ->
  (forall x, inc x (domain f) -> is_cardinal (V x f)) ->
  (forall x, inc x (domain f) -> V x f <> 0c) ->
  inc j (domain f) -> (V j f) <=c (cardinal_prod f).
Proof.
move=> f j fgf alc alnz jd.
have sjd: (sub (singleton j) (domain f)).
  by move=> t td; awi td; rewrite td.
move: (product_increasing1 fgf alc alnz sjd).
rewrite trivial_cardinal_prod3 //.
by rewrite (cardinal_of_cardinal (alc _ jd)).
Qed.

Definition of a finite family of integers

Definition finite_int_fam f:= fgraph f &
  (forall i, inc i (domain f) -> inc (V i f) Bnat) &
  finite_set (domain f).

A finite sum or product of integers is an integer
Lemma finite_sum_finite_aux: forall f x, finite_int_fam f ->
  sub x (domain f) -> inc (cardinal_sum (restr f x)) Bnat.
Proof.
move=> f x [fgf [alB fsd]] sxd.
move: (sub_finite_set sxd fsd) => fsx.
move: x fsx sxd.
apply: finite_set_induction0.
  move=> _;rewrite trivial_cardinal_sum; fprops.
  rewrite restr_domain1=>//; apply emptyset_sub_any.
move=> a b ap nba st.
move: (alB _ (st _ (inc_tack_on_x b a))) => vb.
rewrite induction_sum0 =>//.
apply BS_plus =>//; apply ap; apply sub_trans with (tack_on a b)=>//.
fprops.
Qed.

Lemma finite_product_finite_aux: forall f x, finite_int_fam f ->
  sub x (domain f) -> inc (cardinal_prod (restr f x)) Bnat.
Proof.
move=> f x [fgf [alB fsd]] sxd.
move: (sub_finite_set sxd fsd) => fsx.
move: x fsx sxd.
apply :finite_set_induction0.
  rewrite trivial_cardinal_prod;fprops.
  rewrite restr_domain1=>//; apply emptyset_sub_any.
move=> a b ap nba st.
move: (alB _ (st _ (inc_tack_on_x b a))) => vb.
rewrite induction_prod0 =>//.
apply BS_mult =>//; apply ap; apply sub_trans with (tack_on a b)=>//.
fprops.
Qed.

Theorem finite_sum_finite: forall f, finite_int_fam f ->
  inc (cardinal_sum f) Bnat.
Proof.
move=> f fif.
have <- :restr f (domain f) = f.
  move: fif => [fgf _]; apply restr_to_domain; fprops.
apply (finite_sum_finite_aux fif) ; fprops.
Qed.

Theorem finite_product_finite: forall f, finite_int_fam f ->
  inc (cardinal_prod f) Bnat.
Proof.
move=> f fif.
have <-:restr f (domain f)=f.
  move: fif => [fgf _]; apply restr_to_domain; fprops.
apply finite_product_finite_aux; fprops.
Qed.

Finite unions and products of finite sets are finite sets
Lemma finite_union_finite: forall f, fgraph f ->
  (forall i, inc i (domain f) -> finite_set (V i f))
  -> finite_set (domain f) -> finite_set(unionb f).
Proof.
move=> f fgf alf fsd.
set (g:= L (domain f) (fun a => cardinal (V a f))).
have dg: (domain g = domain f) by rewrite /g; bw.
have fif: (finite_int_fam g).
  red; ee; last by ue.
    rewrite /g;gprops.
  by rewrite dg /g; move=> i idf; bw; apply alf.
move: (finite_sum_finite fif); bw => f1.
move:(cardinal_sum_pr1 fgf); rewrite -/g => f2.
apply (le_int_is_int f1 f2).
Qed.

Lemma finite_product_finite_set: forall f, fgraph f ->
  (forall i, inc i (domain f) -> finite_set (V i f))
  -> finite_set (domain f) -> finite_set(productb f).
Proof.
move=> f fgf alf fsd.
set (g:= L (domain f) (fun a => cardinal (V a f))).
have dg: (domain g = domain f) by rewrite /g;bw.
red;rewrite cardinal_prod_pr//- inc_Bnat.
apply finite_product_finite; red; ee; last by ue.
by bw; move=> i idf; bw;apply alf.
Qed.

EIII-5-2 Strict inequalities between integers

If a<b there is a strictly positive c such that c=a*b

Lemma strict_pos_pr: forall a, inc a Bnat ->
  (0c <> a) = (0c <c a).
Proof.
rewrite /cardinal_lt; move=> a aB; apply iff_eq.
  by move=> na; split=>//; apply zero_smallest; fprops.
by move=> [_].
Qed.

Lemma strict_pos_pr1: forall a, inc a Bnat ->
  (a <> 0c) = (0c <c a).
Proof.
move => a aB; rewrite -strict_pos_pr//; apply iff_eq; intuition.
Qed.

Theorem cardinal_lt_pr: forall a b, inc a Bnat -> inc b Bnat ->
  (a <c b = exists c, inc c Bnat & c <> 0c & a +c c = b).
Proof.
move=> a b aB bB; apply iff_eq.
  move=> [ab nab]; move: (cardinal_complement2 bB ab).
  move: (cardinal_complement1 ab).
  set c := (cardinal (complement b a)) => abc cB; exists c;ee.
  dneg cz; move: abc; rewrite cz; aw; fprops.
move=> [c [cB]]; move=> [cnz]; move => <-.
move: (is_lt_succ aB); rewrite (Bsucc_rw aB) => asa.
have p2: (a +c 1c) <=c (a +c c).
   apply sum_increasing2; fprops; apply one_small_cardinal; fprops.
co_tac.
Qed.

Compatibility of sum and product with strict order

Lemma finite_sum2_lt: forall a b a' b',
  inc a Bnat -> inc b Bnat -> inc a' Bnat -> inc b' Bnat->
  a <=c a' -> b <c b' -> (a +c b) <c (a' +c b').
Proof.
move=> a b a' b' aB bB a'B b'B aa'.
rewrite (cardinal_lt_pr bB b'B); move=> [c [cB [cnz]]] <-.
apply cardinal_le_lt_trans with (a' +c b).
apply sum_increasing2; fprops.
rewrite card_plusA cardinal_lt_pr; hprops; exists c; ee.
Qed.

Lemma finite_sum3_lt: forall a a' b,
  inc a Bnat -> inc b Bnat -> inc a' Bnat ->
  a <c a' -> (a +c b) <c (a'+c b).
Proof.
move=> a a' b aB bB a'B aa'.
rewrite card_plusC (card_plusC a' b).
apply finite_sum2_lt; fprops.
Qed.

Lemma finite_prod2_lt: forall a b a' b',
  inc a Bnat -> inc b Bnat -> inc a' Bnat -> inc b' Bnat->
  a <=c a' -> b <c b' -> a' <> 0c ->
  (a *c b) <c (a' *c b').
Proof.
move=> a b a' b' aB bB a'B b'B aa'.
rewrite (cardinal_lt_pr bB b'B); move=> [c [cB [cnz]]] <- anz.
apply cardinal_le_lt_trans with (card_mult a' b).
apply product_increasing2; fprops.
rewrite cardinal_distrib_prod_sum3.
rewrite cardinal_lt_pr; hprops; exists (card_mult a' c); ee.
apply zero_cardinal_product2=>//.
Qed.

Theorem finite_sum_lt: forall f g,
  finite_int_fam f -> finite_int_fam g -> domain f = domain g ->
  (forall i, inc i (domain f) -> (V i f) <=c (V i g)) ->
  (exists i, inc i (domain f) & (V i f) <c (V i g)) ->
  (cardinal_sum f) <c (cardinal_sum g).
Proof.
move=> f g fif fig df ale [i [ifg lti]].
move: (fif)(fig)=> [f1 [f2 f3]] [g1 [g2 g3]].
have idg: inc i (domain g) by rewrite -df.
move: (tack_on_complement ifg) => dtc.
have incd: ~ (inc i (complement (domain f) (singleton i))).
  by srw; move=> [_ aux]; awi aux; elim aux.
have sd: sub (complement (domain f) (singleton i)) (domain f).
  by apply sub_complement.
have sdg: sub (complement (domain g) (singleton i)) (domain g).
  by apply sub_complement.
rewrite (induction_sum1 f1 dtc incd).
rewrite df in dtc incd; rewrite (induction_sum1 g1 dtc incd).
have afB: inc (V i f) Bnat by apply f2.
have agB: inc (V i g) Bnat by apply g2.
apply finite_sum2_lt =>//.
  apply finite_sum_finite_aux=>//.
  apply finite_sum_finite_aux =>//; ue.
apply sum_increasing;fprops.
  do 2 rewrite restr_domain=>//; ue.
rewrite restr_domain1 //.
move=> x xd; bw=>//; rewrite - ?df;fprops.
Qed.

Theorem finite_product_lt: forall f g,
  finite_int_fam f -> finite_int_fam g -> domain f = domain g ->
  (forall i, inc i (domain f) -> (V i f) <=c (V i g)) ->
  (exists i, inc i (domain f) & (V i f) <c (V i g)) ->
  (forall i, inc i (domain g) -> (V i g) <> 0c) ->
  (cardinal_prod f) <c (cardinal_prod g).
Proof.
move=> f g fif fig df ale [i [ifg lti]] alne.
move: (fif)(fig)=> [f1 [f2 f3]] [g1 [g2 g3]].
move: (partition_complement ifg)=> pfa.
have idg: inc i (domain g) by rewrite -df.
have afB: inc (V i f) Bnat by apply f2.
have agB: inc (V i g) Bnat by apply g2.
have sd: sub (complement (domain f) (singleton i)) (domain f).
  by apply sub_complement.
have sdg: sub (complement (domain g) (singleton i)) (domain g).
  by apply sub_complement.
have incd: ~ (inc i (complement (domain f) (singleton i))).
  by srw; move=> [_ aux]; awi aux; elim aux.
move: (tack_on_complement ifg) => dtc.
rewrite (induction_prod1 f1 dtc incd).
rewrite df in dtc incd; rewrite (induction_prod1 g1 dtc incd).
apply finite_prod2_lt =>//.
      apply finite_product_finite_aux=>//.
    apply finite_product_finite_aux =>//; ue.
  apply product_increasing;fprops.
    do 2 rewrite restr_domain=>//; ue.
  rewrite restr_domain1 //.
  move=> x xd; bw=>//; rewrite - ?df;fprops.
rewrite -zero_cardinal_product; fprops.
rewrite restr_domain1; fprops.
move=> j jdg; move: (jdg); srw.
move=> [jd _]; bw; apply: (alne _ jd).
Qed.

Lemma finite_lt_a_ab: forall a b, inc a Bnat -> inc b Bnat ->
  a <> 0c -> 1c <c b -> a <c (a *c b).
Proof.
move=> a b aB bB naz c1b.
move: (Bnat_is_cardinal aB) => ca.
have laa: (a <=c a) by fprops.
move: (finite_prod2_lt aB inc1_Bnat aB bB laa c1b naz).
by rewrite one_unit_prodr.
Qed.

Compatibility of power and order
Lemma cardinal_le_a_apowb: forall a b,
  is_cardinal a -> a <> 0c -> 1c <=c b -> a <=c (a ^c b).
Proof.
move=> a b ca nza l1b.
have laa: (a <=c a) by fprops.
move: (power_increasing1 nza laa l1b).
by rewrite power_x_1.
Qed.

Lemma non_zero_apowb: forall a b,
  is_cardinal a -> is_cardinal b ->
  a<> 0c -> (a ^c b) <> 0c.
Proof.
move=> a b ca cb anz.
rewrite -(one_small_cardinal2 ca) in anz.
have lbb: (b <=c b) by fprops.
move: (power_increasing1 card_one_not_zero anz lbb).
rewrite power_1_x one_small_cardinal2; fprops.
rewrite /card_pow; fprops.
Qed.

Lemma finite_power_lt1: forall a a' b,
  inc a Bnat -> inc a' Bnat -> inc b Bnat->
  a <c a' -> b <> 0c -> (a ^c b) <c (a' ^c b).
Proof.
move=> a a' b aB a'B bB aa' nzb.
move: (predc_pr bB nzb) => [pB bs].
have nza': a'<> 0c.
  by move => h; rewrite h in aa'; apply (zero_smallest1 aa').
case (equal_or_not a 0c).
  move=> ->; rewrite power_0_x //; split.
   move: (BS_pow a'B bB) => poB; apply zero_smallest; fprops.
  have: a' ^c b <> 0c by apply non_zero_apowb;fprops.
  intuition.
move => nz.
rewrite bs (pow_succ a pB)(pow_succ a' pB).
have p1: a <=c a' by move: aa' => [ok _].
have p2: (predc b) <=c (predc b) by fprops.
move: (BS_pow a'B pB) => ca'.
move: (power_increasing1 nz p1 p2) => p3.
apply (finite_prod2_lt (BS_pow aB pB) aB ca' a'B p3 aa').
apply non_zero_apowb;fprops.
Qed.

Lemma finite_power_lt2: forall a b b',
  inc a Bnat -> inc b Bnat -> inc b' Bnat->
  b <c b' -> 1c <c a -> (a ^c b) <c (a ^c b').
Proof.
move=> a b b' aB bB b'B bb' l1a.
have anz: a <> 0c.
  move: l1a=> [h _]; rewrite- one_small_cardinal2; fprops.
move: bb'; rewrite (cardinal_lt_pr bB b'B);move=> [c [cB [cnz]]] <-.
rewrite power_of_sum2.
apply finite_lt_a_ab; hprops.
apply non_zero_apowb; fprops.
rewrite - one_small_cardinal2 in cnz; fprops.
have ca:is_cardinal a by fprops.
move :(cardinal_le_a_apowb ca anz cnz)=> aux; co_tac.
Qed.

Lemma lt_a_power_b_a: forall a b, inc a Bnat -> inc b Bnat ->
  1c <c b -> a <c (b ^c a).
Proof.
move=> a b aB bB l1b; move: a aB; apply cardinal_c_induction.
  rewrite power_x_0; apply zero_lt_one.
move=> n nB.
move: (BS_pow bB nB) => pB.
have cb: is_cardinal b by fprops.
rewrite -(lt_n_succ_le nB pB) => le1.
apply cardinal_le_lt_trans with (b ^c n) =>//.
rewrite pow_succ //;apply (finite_lt_a_ab pB bB) => //.
apply non_zero_apowb; fprops.
move: l1b=> [l1b _]; rewrite -one_small_cardinal2 //; fprops.
Qed.

Injectivity of sum and product
Lemma plus_simplifiable_left: forall a b b',
  inc a Bnat -> inc b Bnat -> inc b' Bnat ->
  a +c b = a +c b' -> b = b'.
Proof.
move=> a b b' aB bB bB' eql.
have laa: a <=c a by fprops.
case (Bnat_total_order1 bB bB') =>//; case => aux.
move: (finite_sum2_lt aB bB aB bB' laa aux)=> [_ h]; contradiction.
move: (finite_sum2_lt aB bB' aB bB laa aux)=> [_ h].
symmetry in eql;contradiction.
Qed.

Lemma plus_simplifiable_right: forall a b b',
  inc a Bnat -> inc b Bnat -> inc b' Bnat ->
  b +c a = b' +c a -> b = b'.
Proof.
move=> a b b' aB bB b'B;
rewrite card_plusC (card_plusC b' a).
by apply plus_simplifiable_left.
Qed.

Lemma mult_simplifiable_left: forall a b b',
  inc a Bnat -> inc b Bnat -> inc b' Bnat ->
  a <> 0c -> a *c b = a *c b' -> b = b'.
Proof.
move=> a b b' aB bB bB' naz eql.
have laa: a <=c a by fprops.
case (Bnat_total_order1 bB bB') =>//; case => aux.
move: (finite_prod2_lt aB bB aB bB' laa aux naz)=> [_ h]; contradiction.
move: (finite_prod2_lt aB bB' aB bB laa aux naz)=> [_ h].
symmetry in eql;contradiction.
Qed.

Lemma mult_simplifiable_right: forall a b b',
  inc a Bnat -> inc b Bnat -> inc b' Bnat ->
  a <> 0c -> b *c a = b' *c a -> b = b'.
Proof.
move=> a b b' aB bB b'B;
rewrite card_multC (card_multC b' a).
by apply mult_simplifiable_left.
Qed.

cardinal difference

Definition card_sub a b := cardinal (complement a b).
Notation "x -c y" := (card_sub x y) (at level 50).

Lemma card_sub_wrong: forall a b, inc a Bnat -> inc b Bnat ->
  ~ (b <=c a) -> a -c b = 0c.
Proof.
move=> a b aB bB p.
case (Bnat_total_order2 bB aB); first by move=> np; contradiction.
rewrite /cardinal_lt/cardinal_le; move => [ [_ [_ ab]] nab].
rewrite /card_sub.
have -> : complement a b = emptyset.
  by apply is_emptyset => t; srw; move=> [ta ntb]; elim ntb; apply ab.
rewrite cardinal_emptyset; fprops.
Qed.

Lemma BS_sub: forall a b, inc a Bnat -> inc b Bnat ->
  inc (a -c b) Bnat.
Proof.
move=> a b aB bB; case (p_or_not_p (b <=c a)).
  apply cardinal_complement2 => //.
move => h; rewrite card_sub_wrong //; fprops.
Qed.

Hint Resolve BS_sub: hprops.

Lemma card_sub_pr: forall a b,
  b <=c a -> b +c (a -c b) = a.
Proof. apply cardinal_complement1. Qed.

Lemma card_sub_pr0: forall a b, inc a Bnat-> inc b Bnat ->
  b <=c a -> (inc (a -c b) Bnat & b +c (a -c b) = a).
Proof. move=> a b aB bB h; split; [ hprops | by apply card_sub_pr]. Qed.

Lemma card_sub_pr1: forall a b, inc a Bnat-> inc b Bnat ->
  (a +c b) -c b = a.
Proof.
move=> a b aB bB.
move: (BS_plus aB bB)=> sB.
move: (Bsum_increasing3 bB aB).
rewrite card_plusC=> aux.
move: (card_sub_pr aux); rewrite card_plusC=> aux'.
apply (plus_simplifiable_right bB (BS_sub sB bB) aB aux').
Qed.

Lemma card_sub_rpr: forall a b,
  b <=c a -> (a -c b) +c b = a.
Proof. by move=> a b leaB;rewrite card_plusC; apply card_sub_pr. Qed.

Lemma card_sub_pr2: forall a b c, inc a Bnat-> inc b Bnat ->
  a +c b = c -> c -c b = a.
Proof. by move=> a b c aB bB h; move: (card_sub_pr1 aB bB); rewrite h. Qed.

Lemma minus_n_nC: forall a, a -c a = 0c.
Proof.
by move=> a; rewrite /card_sub complement_itself cardinal_emptyset.
Qed.

Lemma minus_n_0C: forall a, inc a Bnat -> a -c 0c = a.
Proof. move=> a aB; apply card_sub_pr2; fprops; aw; fprops.
Qed.

Lemma card_sub_pr4: forall a b a' b', inc a Bnat-> inc b Bnat ->
  inc a' Bnat-> inc b' Bnat ->
  a <=c b -> a' <=c b' ->
  (b +c b') -c (a +c a') = (b -c a) +c (b' -c a').
Proof.
move=> a b a' b' aB bB a'B b'B ab a'b'.
apply card_sub_pr2; [ hprops | hprops |].
have aux: ((b -c a) +c b') +c a = b' +c b.
  by rewrite (card_plusC _ b') - card_plusA card_sub_rpr.
by rewrite (card_plusC a a') card_plusA - (card_plusA _ _ a')
 (card_sub_rpr a'b') aux card_plusC.
Qed.

Lemma card_sub_associative: forall a b c,
  inc a Bnat -> inc b Bnat -> inc c Bnat ->
  (b +c c) <=c a -> (a -c b) -c c = a -c (b +c c).
Proof.
move=> a b c aB bB cB h.
move: (card_sub_pr h) => aux.
apply card_sub_pr2 =>//; first by hprops.
symmetry; apply card_sub_pr2; hprops.
by rewrite - card_plusA card_plusC (card_plusC c b).
Qed.

Lemma prec_pr1: forall a, inc a Bnat ->
  predc a = a -c 1c.
Proof.
move=> a aB.
case (equal_or_not a 0c) => anz.
   rewrite /predc /card_sub anz /card_zero union_empty.
   have -> : complement emptyset card_one = emptyset.
   apply is_emptyset => t; srw; move=> [te]; empty_tac0.
   rewrite cardinal_emptyset //.
have le1a: (1c <=c a) by apply one_small_cardinal; fprops.
move: (card_sub_pr le1a) => h.
move: (BS_sub aB inc1_Bnat) => fa1.
move: (predc_pr aB anz)=> [_ bb].
have aux: (succ (a -c 1c) = a).
  by rewrite (Bsucc_rw fa1) card_plusC.
rewrite - {1} aux;apply predc_pr2; fprops.
Qed.

Lemma card_sub_non_zero1: forall a b, inc a Bnat ->
  b <c a -> a -c b <> 0c.
Proof. apply cardinal_complement3 => //. Qed.

Lemma card_sub_non_zero: forall a b, inc a Bnat -> inc b Bnat ->
  (succ b) <=c a -> a -c b <> 0c.
Proof.
move=> a b aB bB lesba.
apply cardinal_complement3 => //; rewrite - lt_n_succ_le //.
Qed.

Lemma card_sub_associative1: forall a b, inc a Bnat -> inc b Bnat ->
  (succ b) <=c a -> predc (a -c b) = a -c (succ b).
Proof.
move=> a b aB bB; rewrite (Bsucc_rw bB) => h'.
rewrite - card_sub_associative; fprops.
apply prec_pr1; hprops.
Qed.

Lemma sub_le_symmetry: forall a b,
  b <=c a -> (a -c b) <=c a.
Proof.
move=> a b ba; move:(card_sub_pr ba) => aux.
rewrite -{2} aux card_plusC; apply sum_increasing3.
  rewrite /card_sub; fprops.
co_tac.
Qed.

Lemma sub_lt_symmetry: forall n p, inc p Bnat ->
  n <c p -> (predc (p -c n)) <c p.
Proof.
move=> n p pB ltnp.
have nB: (inc n Bnat) by Bnat_tac.
move: (card_sub_non_zero1 pB ltnp) => snz.
move: ltnp => [np nnp].
move: (card_sub_pr0 pB nB np) => [sf aux].
move: (predc_pr sf snz)=> [aa bb].
rewrite -lt_n_succ_le; fprops.
rewrite -bb;apply sub_le_symmetry =>//.
Qed.

Lemma double_sub: forall n p, inc n Bnat ->
  p <=c n -> n -c (n -c p) = p.
Proof.
move=> n p nB lepn.
move: (le_int_in_Bnat lepn nB) => pB.
apply card_sub_pr2; [exact pB | hprops | apply card_sub_pr=>//].
Qed.

Lemma succ_sub: forall a b, inc a Bnat -> inc b Bnat ->
  (succ b) <=c a -> a -c b = succ (a -c (succ b)).
Proof.
move=> a b aB bB aux; move: (BS_succ bB) => sB.
apply card_sub_pr2 ; hprops.
have sb: inc (a -c (succ b)) Bnat by hprops.
rewrite (Bsucc_rw sb) - card_plusA.
have->: (1c +c b = succ b) by rewrite (Bsucc_rw bB) card_plusC.
apply card_sub_rpr =>//.
Qed.


Lemma Bnat_le_reflexive: forall a, inc a Bnat -> a <=N a.
Proof. move => a aB; red; ee. Qed.

Lemma Bnat_le_transitive: forall a b c, a <=N b -> b <=N c -> a <=N c.
Proof.
move => a b c; rewrite - !Bnat_order_le.
move => ab bc; move: Bnat_order_worder=> [or wor]; order_tac.
Qed.

Lemma Bnat_le_antisymmetric: forall a b, a <=N b -> b <=N a -> a = b.
Proof.
move=> a b; rewrite - !Bnat_order_le.
move => ab ba; move: Bnat_order_worder=> [or wor]; order_tac.
Qed.

Lemma Bnat_total_order: forall a b, inc a Bnat -> inc b Bnat ->
  a <=N b \/ b <N a.
Proof.
move=> a b aB bB; rewrite /Bnat_lt - ! Bnat_order_le.
move: (worder_total Bnat_order_worder)=> [or prop].
case (equal_or_not b a).
  move=>->; left; order_tac; by rewrite Bnat_order_substrate.
rewrite Bnat_order_substrate in prop; case (prop _ _ aB bB); auto.
Qed.

Lemma Bnat_zero_smallest: forall a, inc a Bnat -> 0c <=N a.
Proof. move=> a aB; red; ee; fprops. Qed.

Lemma Bnat_zero_smallest1: forall a, a <=N 0c -> a = 0c.
Proof.
move=> a h; apply Bnat_le_antisymmetric=>//.
move: h; rewrite /Bnat_le; move => [aB _ ]; ee.
Qed.

Lemma plus_le_simplifiable: forall a b c,
  inc a Bnat -> inc b Bnat -> inc c Bnat->
  (a +c b) <=c (a +c c) -> b <=c c.
Proof.
move=>a b c aB bB cB abc.
case (Bnat_total_order2 bB cB) => // ltcb.
have aa: cardinal_le a a by fprops.
move: (finite_sum2_lt aB cB aB bB aa ltcb)=> h; co_tac.
Qed.

Lemma plus_lt_simplifiable: forall a b c,
  inc a Bnat -> inc b Bnat -> inc c Bnat->
  (a +c b) <c (a +c c) -> b <c c.
Proof.
move=> a b c aB bB cB [abc nac]; split.
apply: (plus_le_simplifiable aB)=>//.
by dneg bc; rewrite bc.
Qed.

Lemma mult_le_simplifiable: forall a b c,
  inc a Bnat -> inc b Bnat -> inc c Bnat-> a <> 0c ->
  (a *c b) <=c (a *c c) -> b <=c c.
Proof.
move=>a b c aB bB cB naz abc.
case (Bnat_total_order2 bB cB) => // ltcb.
have aa: cardinal_le a a by fprops.
move: (finite_prod2_lt aB cB aB bB aa ltcb naz)=> h; co_tac.
Qed.

Lemma mult_lt_simplifiable: forall a b c,
  inc a Bnat -> inc b Bnat -> inc c Bnat-> a <> 0c ->
  (a *c b) <c (a *c c) -> b <c c.
Proof.
move=> a b c aB bB cB naz [abc nac]; split.
apply: (mult_le_simplifiable aB)=>//.
by dneg bc; rewrite bc.
Qed.

Lemma card_sub_pr5: forall a b c, inc a Bnat -> inc b Bnat -> inc c Bnat ->
  (a +c c) -c (b +c c) = a -c b.
Proof.
move => a b c aB bB cB.
case (p_or_not_p (b <=c a)) => ab.
  apply card_sub_pr2 ; [hprops | hprops |].
  rewrite card_plusA (card_plusC _ b) card_sub_pr//.
symmetry; rewrite card_sub_wrong//card_sub_wrong//; hprops.
rewrite card_plusC (card_plusC a c).
dneg ba.
apply (plus_le_simplifiable cB bB aB ba).
Qed.

Lemma card_sub_pr6: forall a b, inc a Bnat -> inc b Bnat ->
  (succ a) -c (succ b) = a -c b.
Proof.
move=> a b aB bB; rewrite !Bsucc_rw //;apply card_sub_pr5 =>//; fprops.
Qed.

EIII-5-3 Intervals in sets of integers

Definition interval_Bnat a b := interval_cc Bnat_order a b.
Definition interval_co_0a a:= interval_co Bnat_order 0c a.
Definition interval_cc_0a a:= interval_Bnat 0c a.
Definition interval_cc_1a a:= interval_Bnat 1c a.

Definition interval_Bnato a b :=
  graph_on cardinal_le (interval_cc Bnat_order a b).
Definition interval_Bnatco a :=
  graph_on cardinal_le (interval_co_0a a).

Lemma interval_Bnat_pr: forall a b x, inc a Bnat -> inc b Bnat ->
  inc x (interval_Bnat a b) = (a <=c x & x <=c b).
Proof.
rewrite /interval_Bnat=> a b x aB bB.
rewrite Bnat_interval_cc_pr // /Bnat_le; apply iff_eq; first by intuition.
move=> [leax lexb]; have xB: inc x Bnat by Bnat_tac.
ee.
Qed.

Hint Rewrite interval_Bnat_pr : sw.

Lemma interval_Bnat_pr0: forall b x, inc b Bnat ->
  inc x (interval_cc_0a b) = x <=c b.
Proof.
rewrite /interval_cc_0a;move=> b x bB; srw; fprops.
apply iff_eq; first by intuition.
move=> xb; ee.
move: xb => [cx _];apply zero_smallest=> //.
Qed.

Lemma interval_Bnat_pr1: forall b x, inc b Bnat ->
  inc x (interval_cc_1a b) = (x <> 0c & x <=c b).
Proof.
rewrite/interval_cc_1a; move:inc1_Bnat => oB a b aB ; srw.
apply iff_eq.
  move=> [ob ba]; ee.
  move=> bz; rewrite bz in ob.
  move: (zero_smallest2 ob); apply card_one_not_zero.
move=> [bz cle]; ee.
move: cle => [cb _];apply one_small_cardinal=> //.
Qed.

Lemma interval_Bnat_pr1b: forall b x, inc b Bnat ->
  inc x (interval_cc_1a b) = (1c <=c x & x <=c b).
Proof.
move=> a b aB; move: inc1_Bnat => oB.
rewrite (interval_Bnat_pr _ oB aB)//.
Qed.

Lemma sub_interval_Bnat: forall a b,
  sub (interval_Bnat a b) Bnat.
Proof.
move =>a b; rewrite /interval_Bnat /interval_cc Bnat_order_substrate.
apply Z_sub.
Qed.

Lemma sub_interval_co_0a_Bnat: forall a, sub (interval_co_0a a) Bnat.
Proof.
move=> a; rewrite /interval_co_0a /interval_co Bnat_order_substrate.
apply Z_sub.
Qed.

Lemma interval_co_0a_pr2: forall a x, inc a Bnat ->
  inc x (interval_co_0a a) = x <c a.
Proof.
move=> a x aB; rewrite /interval_co_0a Bnat_interval_co_pr1; fprops.
apply iff_eq;first by move=> [_].
by move=> h; ee; apply zero_smallest; move: h => [[h _ ] _].
Qed.

Hint Rewrite interval_co_0a_pr2 interval_Bnat_pr0: sw.

Lemma interval_co_0a_pr3: forall a x, inc a Bnat ->
  inc x (interval_co_0a (succ a)) = x <=c a.
Proof. move=> a x aB; srw; hprops. Qed.

Lemma interval_co_cc: forall p, inc p Bnat ->
  interval_cc_0a p = interval_co_0a (succ p).
Proof.
move=>p pB;set_extens x; rewrite interval_co_0a_pr3 //interval_Bnat_pr0 //.
Qed.

Lemma interval_cc_0a_increasing: forall a b, inc b Bnat ->
  a <=c b ->
  sub (interval_cc_0a a) (interval_cc_0a b).
Proof.
move=> a b bB ab.
have aB: inc a Bnat by Bnat_tac.
move=> t; srw; move=> ta;co_tac.
Qed.

Lemma interval_cc_0a_increasing1: forall a, inc a Bnat ->
  sub (interval_cc_0a a) (interval_cc_0a (succ a)).
Proof.
move=> a aB;apply interval_cc_0a_increasing; hprops.
apply is_le_succ; fprops.
Qed.

Lemma inc_a_interval_co_succ: forall a, inc a Bnat ->
  inc a (interval_co_0a (succ a)).
Proof. move=> a aB; srw; fprops; hprops. Qed.

Lemma interval_co_0a_increasing: forall a, inc a Bnat ->
  sub (interval_co_0a a) (interval_co_0a (succ a)).
Proof. move=> a aB t;srw; hprops; by move=> [h _]. Qed.

Lemma interval_co_0a_increasing1: forall a b, inc a Bnat -> inc b Bnat ->
  a <=c b -> sub (interval_co_0a a) (interval_co_0a b).
Proof.
move=> a b aB bB ab r; rewrite ! Bnat_interval_co_pr1; fprops.
move=> [h1 h2]; split=> //; co_tac.
Qed.

Lemma interval_co_pr4: forall n, inc n Bnat ->
  ( (tack_on (interval_co_0a n) n = (interval_co_0a (succ n)))
    & ~(inc n (interval_co_0a n))).
Proof.
move=> n nB.
move: (BS_succ nB) => snB.
have fn:finite_c n by rewrite - inc_Bnat.
split; last by srw; rewrite /cardinal_lt; intuition.
set_extens x; rewrite tack_on_rw !interval_co_0a_pr2 //; srw;
 rewrite/cardinal_lt.
    case; ee; move=> -> ; fprops.
case (equal_or_not x n); intuition.
Qed.

Lemma interval_bn_pr5: forall n, inc n Bnat ->
  let si := (interval_Bnat 1c n) in
  ( (tack_on si 0c = interval_cc_0a n) & ~(inc 0c si)).
Proof.
move=> n nB si.
move: inc1_Bnat => oB.
split; rewrite /si.
   set_extens x; aw; srw.
     case; first by intuition.
     move=> ->; fprops.
   move => h; case (equal_or_not x 0c) => xz; [right => //| left ;ee ].
   apply one_small_cardinal => //; co_tac.
rewrite interval_Bnat_pr1//; intuition.
Qed.

Lemma cardinal_c_induction5: forall (r:Set -> Prop) a,
  inc a Bnat -> r 0c ->
  (forall n, n <c a -> r n -> r (succ n))
  -> (forall n, n <=c a -> r n).
Proof.
move=> r a aB r0 rs.
have rss: forall n,
  inc n (interval_co_0a a) -> r n -> r (succ n).
  by move=> n; srw; apply rs.
move: (cardinal_c_induction3_v (r:=r) inc0_Bnat aB r0 rss) => h n na.
by apply h; rewrite interval_Bnat_pr0.
Qed.

Lemma interval_Bnato_worder: forall a b, inc a Bnat -> inc b Bnat ->
  worder (interval_Bnato a b).
Proof.
move=> a b aB bB ; rewrite /interval_Bnato.
move: wordering_cardinal_le => wo.
have r: (forall x, inc x (interval_cc Bnat_order a b) -> x <=c x).
 move=> x; srw; rewrite /Bnat_le; move => [[_ [xB _]] _]; fprops.
by move: (wordering_pr wo r)=> [_].
Qed.

Lemma interval_Bnato_substrate: forall a b, inc a Bnat -> inc b Bnat ->
  substrate (interval_Bnato a b) = interval_Bnat a b.
Proof.
move=> a b aB bB ; rewrite /interval_Bnato.
move: wordering_cardinal_le => wo.
have r: (forall x, inc x (interval_cc Bnat_order a b) -> x <=c x).
 move=> x; srw; rewrite /Bnat_le; move => [[_ [xB _]] _]; fprops.
by move: (wordering_pr wo r)=> [].
Qed.

Lemma interval_Bnato_related: forall a b x y, inc a Bnat -> inc b Bnat ->
  gle (interval_Bnato a b) x y = (inc x (interval_Bnat a b) &
    inc y (interval_Bnat a b) &x <=c y).
Proof.
move=> a b x y aB bB; rewrite /interval_Bnato; aw.
by rewrite /gle graph_on_rw1.
Qed.

Lemma interval_Bnato_related1: forall a b x y, inc a Bnat -> inc b Bnat ->
  gle (interval_Bnato a b) x y =
  (a <=c x & a <=c y & x <=c b & y <=c b & x <=c y).
Proof.
move=> a b x y aB bB; rewrite interval_Bnato_related //; srw.
rewrite /Bnat_le; apply iff_eq; first by intuition.
move=> [ax [ay [xb [yb xy]]]]; ee.
Qed.

Lemma interval_Bnato_related2: forall a b x y, inc a Bnat -> inc b Bnat ->
  gle (interval_Bnato a b) x y =
   (a <=c x & y <=c b & x <=c y).
Proof.
move=> a b x y aB bB; rewrite interval_Bnato_related1 //.
apply iff_eq; [intuition | move=> [ax [yb xy]]; ee; co_tac ].
Qed.

Lemma interval_Bnatco_worder: forall a, inc a Bnat ->
  worder (interval_Bnatco a).
Proof.
rewrite /interval_Bnatco=> a aB.
move: wordering_cardinal_le => wo.
have r: forall x, inc x (interval_co_0a a) -> x <=c x.
  move=> x xi; move: (sub_interval_co_0a_Bnat xi) => xB; fprops.
by move: (wordering_pr wo r)=> [].
Qed.

Lemma interval_Bnatco_substrate: forall a, inc a Bnat ->
  substrate (interval_Bnatco a) = interval_co_0a a.
Proof.
rewrite /interval_Bnatco=> a aB.
move: wordering_cardinal_le => wo.
have r: forall x, inc x (interval_co_0a a) -> x <=c x.
  move=> x xi; move: (sub_interval_co_0a_Bnat xi) => xB; fprops.
by move: (wordering_pr wo r)=> [].
Qed.

Lemma interval_Bnatco_related: forall a x y, inc a Bnat ->
  gle (interval_Bnatco a) x y = (x <=c y & y <c a).
Proof.
move=> a x y aB; rewrite /interval_Bnatco /gle graph_on_rw1; srw.
apply iff_eq; ee.
move=> [xy ya]; ee; co_tac.
Qed.

Lemma segment_Bnat_order: forall x, inc x Bnat ->
  segment Bnat_order x = interval_co Bnat_order 0c x.
Proof.
move=> x xB; rewrite /segment /interval_co /interval_uo /glt.
set_extens t; rewrite ! Z_rw; ee.
move=> [ts [tx ntx]]; ee.
move: tx; rewrite ! Bnat_order_le; rewrite /Bnat_le; ee.
Qed.

Definition rest_plus_interval a b :=
  BL(fun z => z +c b)(interval_Bnat 0c a)
    (interval_Bnat b (a +c b)).

Definition rest_minus_interval a b :=
  BL(fun z => z -c b) (interval_Bnat b (a +c b))
  (interval_Bnat 0c a).

Lemma sub_increasing2: forall a b c, inc a Bnat -> inc b Bnat -> inc c Bnat->
  c <=c (a +c b) -> (c -c b) <=c a.
Proof.
move=> a b c aB bB cB cab.
case (p_or_not_p (b <=c c)) => h.
  apply (plus_le_simplifiable (a:= b)); hprops.
   rewrite card_plusC in cab; rewrite card_sub_pr //.
rewrite card_sub_wrong //; fprops.
Qed.

Theorem restr_plus_interval_isomorphism: forall a b, inc a Bnat -> inc b Bnat->
 (bijection (rest_plus_interval a b)
 & bijection (rest_minus_interval a b)
 & (rest_minus_interval a b) = inverse_fun (rest_plus_interval a b)
 &
   order_isomorphism (rest_plus_interval a b)
  (interval_Bnato 0c a)
  (interval_Bnato b (a +c b))).
Proof.
move=> a b aB bB.
move:(inc0_Bnat) => zB.
move: (BS_plus aB bB) => abB.
set E1:= interval_cc_0a a.
set E2:= interval_Bnat b (a +c b).
have tap: transf_axioms (fun z => z +c b) E1 E2.
  rewrite /E1 /E2; move=> z; srw => za; split.
    rewrite card_plusC; apply sum_increasing3; fprops.
    by move: za => [ok _].
  apply sum_increasing2=>//; fprops.
have tam: transf_axioms (fun z => z -c b) E2 E1.
  rewrite /E1 /E2;move=> t; srw; move=> [bt tab].
  have tB: inc t Bnat by Bnat_tac.
  split.
    move: (BS_sub tB bB) => sB;apply zero_smallest; fprops.
  apply sub_increasing2 =>//.
set (f:= rest_plus_interval a b).
set (g:= rest_minus_interval a b).
have ff:is_function f by rewrite /f /rest_plus_interval; apply bl_function.
have fg:is_function g by rewrite /g /rest_plus_interval; apply bl_function.
have sf:source f= target g.
  by rewrite /f/g /rest_plus_interval /rest_minus_interval bl_source bl_target.
have sg:source g = target f.
  by rewrite /f/g /rest_plus_interval /rest_minus_interval bl_source bl_target.
have cfg:f \coP g by red; ee.
have cgf: g \coP f by red; ee.
have c1: g \co f = identity (source f).
  apply function_exten; [fct_tac | fprops | aw | symmetry; aw | ].
  rewrite compose_source;move=> x xsf.
  rewrite (identity_W xsf) (compose_W cgf xsf).
  move: xsf; rewrite /f /rest_plus_interval bl_source=> xs.
  rewrite (bl_W tap xs) /g /rest_minus_interval (bl_W tam (tap _ xs)).
  rewrite card_sub_pr1 =>//; apply (sub_interval_Bnat xs).
have c2: f \co g = identity (source g).
  apply function_exten ; [fct_tac | fprops | aw | symmetry; aw | ].
  rewrite compose_source;move=> x xsg.
  rewrite (identity_W xsg) (compose_W cfg xsg).
  move: xsg; rewrite /g /rest_minus_interval bl_source=> xs.
  rewrite (bl_W tam xs).
  move: xs; rewrite /interval_Bnat Z_rw ! Bnat_order_le /Bnat_le.
  move=> [_ [[_ [xB bx]] [_ [sB xab]] ]].
  have aux: inc (x -c b) (interval_Bnat 0c a).
     rewrite interval_Bnat_pr0 //; apply sub_increasing2 =>//.
  rewrite /f /rest_plus_interval (bl_W tap aux).
  rewrite card_plusC card_sub_pr //.
move: (bijective_from_compose cgf cfg c1 c2)=> [bf [bg gif]]; ee.
have sp : E1 = source (rest_plus_interval a b).
  by rewrite /rest_plus_interval; aw.
move: (interval_Bnato_worder zB aB) => [o1 _ ].
move: (interval_Bnato_worder bB abB)=> [o2 _ ].
red; ee.
- rewrite interval_Bnato_substrate /f/rest_plus_interval // bl_source //.
- rewrite interval_Bnato_substrate /f/rest_plus_interval // bl_target //.
- move=> x y; rewrite /f /rest_plus_interval bl_source => xsf ysf.
  rewrite !interval_Bnato_related // (bl_W tap xsf) (bl_W tap ysf).
  apply iff_eq; move=> [_ [ _ xy]].
    split; first (by apply tap); split; first (by apply tap).
    apply sum_increasing2=>//; fprops.
  rewrite card_plusC (card_plusC y b) in xy.
  split; first (by exact); split; first (by exact).
  apply (plus_le_simplifiable bB
      (sub_interval_Bnat xsf)(sub_interval_Bnat ysf) xy).
Qed.

Lemma interval_zero_zero:
  interval_cc_0a 0c = singleton 0c.
Proof.
have zB: inc 0c Bnat by fprops.
set_extens t; aw;srw; first by move => h; rewrite (zero_smallest2 h).
move=> ->; fprops.
Qed.

Lemma cardinal_interval0a: forall a, inc a Bnat ->
  cardinal (interval_cc_0a a) = succ a.
Proof.
apply cardinal_c_induction.
  rewrite succ_zero interval_zero_zero; apply cardinal_singleton.
move => n nB.
move: (BS_succ nB) => snB.
rewrite (interval_co_cc nB) (interval_co_cc snB).
move: (interval_co_pr4 snB) => [res aux]; rewrite - res.
by rewrite cardinal_succ_pr ; [move => -> |].
Qed.

Theorem cardinal_interval: forall a b, a <=N b ->
  cardinal (interval_Bnat a b) = succ (b -c a).
Proof.
move=> a b [aB [bB ab]].
move: (card_sub_pr ab) (BS_sub bB aB) => aux cB.
rewrite card_plusC in aux.
set (c:= b -c a) in *.
move: (restr_plus_interval_isomorphism cB aB) => [b1 _].
have eq: (interval_Bnat 0c c) \Eq (interval_Bnat a b).
  exists (rest_plus_interval c a); rewrite /rest_plus_interval.
  split => //; rewrite bl_source bl_target aux; split => //.
by rewrite -(cardinal_interval0a cB) cardinal_equipotent; eqsym.
Qed.

Lemma finite_set_interval_Bnat: forall a b, a <=N b ->
  finite_set (interval_Bnat a b).
Proof. move=> a b h; red;rewrite cardinal_interval // -inc_Bnat.
move: h => [aB [bB _]]; hprops.
Qed.

Lemma finite_set_interval_co: forall a, inc a Bnat ->
  finite_set (interval_co_0a a).
Proof.
move=> a aB.
have aux:sub (interval_co_0a a) (interval_Bnat 0c a).
  by move=> t; srw; move => [ok _].
have : (finite_set (interval_Bnat 0c a)).
  by apply finite_set_interval_Bnat; apply Bnat_zero_smallest.
apply (sub_finite_set aux).
Qed.

Lemma Bnat_infinite: ~(finite_set Bnat).
Proof.
rewrite /finite_set -inc_Bnat =>h.
move: (sub_smaller (sub_interval_Bnat (a:=0c) (b:=cardinal Bnat))).
rewrite (cardinal_interval0a h).
srw;fprops; rewrite /cardinal_lt; intuition.
Qed.

Lemma cardinal_interval1a: forall a, inc a Bnat ->
  cardinal (interval_cc_1a a) = a.
Proof.
move => a aB.
case (equal_or_not a 0c).
  move=> ->.
  have ->: (interval_cc_1a 0c = emptyset).
    rewrite /interval_cc_1a.
    apply is_emptyset; move=> y; srw; fprops; move=> [c1y cy0].
    move: (cardinal_le_transitive c1y cy0) zero_lt_one => h1 h2.
    co_tac.
 apply cardinal_emptyset.
move => anz.
have aux1: 1c <=c a by apply one_small_cardinal; fprops.
have aux: 1c <=N a by red;ee.
have so: inc (a -c 1c) Bnat by hprops.
rewrite cardinal_interval // Bsucc_rw // card_plusC.
apply card_sub_pr; fprops.
Qed.

Lemma isomorphism_worder_finite: forall r r',
  total_order r -> total_order r' ->
  finite_set (substrate r) -> (substrate r) \Eq (substrate r') ->
  exists_unique (fun f => order_isomorphism f r r').
Proof.
move=> r r' tor tor' fs.
rewrite - cardinal_equipotent=> sc.
have fs': finite_set (substrate r') by red; rewrite -sc.
move:(finite_set_torder_worder tor fs) (finite_set_torder_worder tor' fs').
move => wor wor'; move: (isomorphism_worder wor wor') => iw.
case iw.
  move=> [[f [srf mf]] unq].
  split.
    move: mf => [_ [_ [injf [sf [tf pf]]]]].
    exists f; split;[ fprops | split; [ fprops | split]].
     apply bijective_if_same_finite_c_inj; [ by rewrite -sf -tf | ue| exact].
     ee.
  move =>a b oa ob.
  have pa: (is_segment r' (range (graph a)) & order_morphism a r r').
    move: oa => [_ [or' [[injf sjf] [sf [tf pf]]]]].
    rewrite (surjective_pr3 sjf) -tf.
    split; first (by apply substrate_is_segment =>//); red; ee.
  have pb: (is_segment r' (range (graph b)) & order_morphism b r r').
    move: ob => [_ [or' [[injf sjf] [sf [tf pf]]]]].
    rewrite (surjective_pr3 sjf) -tf.
    split; first (by apply substrate_is_segment =>//); red; ee.
  apply unq =>//.
move=> [[f [srf mf]] unq].
have oi: (order_isomorphism f r' r).
   move: mf => [_ [_ [injf [sf [tf pf]]]]].
   have bf: bijection f by apply bijective_if_same_finite_c_inj;
     rewrite -? sf -?tf;[ symmetry; apply sc| exact| exact].
   red; ee.
move: (inverse_order_isomorphism oi)=> oii.
split; first by exists (inverse_fun f)=>//.
move=> x y ox oy.
suff: (inverse_fun x = inverse_fun y).
  have cx: is_correspondence x by move: ox => [_ [_ [[[[ cf _] _] _] _]]].
  have cy: is_correspondence y by move: oy => [_ [_ [[[[ cf _] _] _] _]]].
  rewrite - {2} (inverse_fun_involutive cx)- {2} (inverse_fun_involutive cy).
  by move=> ->.
apply unq.
   move: (inverse_order_isomorphism ox)=> [_ [or' [[injf sjf] [sf [tf pf]]]]].
    rewrite (surjective_pr3 sjf) -tf.
 split; first (by apply substrate_is_segment =>//); red; ee.
move: (inverse_order_isomorphism oy)=> [_ [or' [[injf sjf] [sf [tf pf]]]]].
rewrite (surjective_pr3 sjf) -tf.
split; first (by apply substrate_is_segment =>//); red; ee.
Qed.

Theorem finite_ordered_interval: forall r, total_order r ->
  finite_set (substrate r) ->
  exists_unique (fun f => order_isomorphism f r
    (interval_Bnato 1c (cardinal (substrate r)))).
Proof.
move=> r tot fs.
move: (fs); rewrite /finite_set -inc_Bnat=> fs'.
move: (worder_total (interval_Bnato_worder inc1_Bnat fs')) => tor.
move: (cardinal_interval1a fs').
rewrite cardinal_equipotent /interval_cc_1a - interval_Bnato_substrate; fprops.
move => h;apply (isomorphism_worder_finite tot tor fs).
by eqsym.
Qed.

Lemma cardinal_interval_co_0a: forall a, inc a Bnat -> a <> 0c ->
  cardinal (interval_cc_0a (predc a)) = a.
Proof.
move=> a aB naz.
move: (predc_pr aB naz) => [fp ass].
have aux: (Bnat_le 0c (predc a)) by red; ee; bw.
by rewrite (cardinal_interval aux) minus_n_0C.
Qed.

Lemma interval_co_0a_pr1: forall a, inc a Bnat -> a <> 0c ->
  interval_Bnat 0c (predc a) = interval_co_0a a.
Proof.
move: inc0_Bnat => zB.
move=> a aB naz;move: (predc_pr aB naz) => [fp ass].
by set_extens x; srw; rewrite -lt_is_le_succ // -ass.
Qed.

Lemma emptyset_interval_00: interval_co_0a 0c = emptyset.
Proof.
apply is_emptyset; move=> y;srw; fprops =>h.
elim (zero_smallest1 h).
Qed.

Lemma cardinal_interval_co_0a1: forall a, inc a Bnat ->
  cardinal (interval_co_0a a) = a.
Proof.
move=> a aB; case (equal_or_not a 0c).
  move => ->; rewrite emptyset_interval_00; apply cardinal_emptyset.
move=> h; rewrite - interval_co_0a_pr1//; apply cardinal_interval_co_0a=>//.
Qed.

Theorem finite_ordered_interval1: forall r, total_order r ->
  finite_set (substrate r) ->
  exists_unique (fun f => order_isomorphism f r
    (interval_Bnatco (cardinal (substrate r)))).
Proof.
move=> r tor fs.
move: (finite_set_torder_worder tor fs)=> wor.
move: (fs); rewrite /finite_set -inc_Bnat => fs'.
move: (cardinal_interval_co_0a1 fs').
rewrite cardinal_equipotent - interval_Bnatco_substrate // => e.
move: (worder_total (interval_Bnatco_worder fs')) => tor'.
by apply (isomorphism_worder_finite tor tor' fs); eqsym.
Qed.

Lemma partition_tack_on_intco: forall a, inc a Bnat ->
  partition_fam (variantLc (interval_co_0a a)
    (singleton a)) (interval_co_0a (succ a)).
Proof.
move=> a aB.
have fa:finite_c a by rewrite - inc_Bnat.
move: (interval_co_pr4 aB) => [e p]; rewrite - e.
by apply partition_tack_on.
Qed.

Lemma interval_co_0a_restr: forall a f, inc a Bnat ->
  L (interval_co_0a a) f
    = (restr (L (interval_co_0a (succ a)) f) (interval_co_0a a)).
Proof.
move=> a f aB.
set aux := (interval_co_0a (succ a)).
have Ha:sub (interval_co_0a a) (domain (L aux f)).
  by bw; apply interval_co_0a_increasing.
have fg: fgraph (L aux f) by gprops.
apply fgraph_exten; rewrite? restr_domain1 // ; bw; gprops.
  apply restr_fgraph; gprops.
move=> x; srw => xa; rewrite restr_ev //; srw.
rewrite /aux;move: (xa) => [ xle xne]; bw; srw; hprops.
Qed.

Lemma partition_tack_on_int: forall n, inc n Bnat ->
  partition_fam (variantLc (interval_cc_1a n) (singleton 0c))
    (interval_cc_0a n).
Proof.
move=> a aB.
have fa:finite_c a by rewrite - inc_Bnat.
move: (interval_bn_pr5 aB) => [e p]. rewrite - e.
by apply partition_tack_on.
Qed.

Lemma interval_int_restr: forall a f, inc a Bnat ->
  L (interval_cc_1a a) f = restr (L (interval_cc_0a a) f) (interval_cc_1a a).
Proof.
move=> a f aB.
move: (interval_bn_pr5 aB)=>/=.
rewrite / interval_cc_1a.
set si:= (interval_Bnat 1c a).
set bi := interval_cc_0a a.
move=> [ta _].
have incl: sub si (domain (L bi f)) by rewrite -ta; bw; fprops.
have g1: fgraph (L bi f) by gprops.
apply fgraph_exten;gprops; bw; gprops.
  apply restr_fgraph =>//.
  rewrite restr_domain1 //.
rewrite L_domain in incl.
by move=> x xsi; bw; apply incl.
Qed.

A finite sum or product can be defined by induction
Lemma induction_on_sum: forall a f, inc a Bnat ->
  let iter := fun n=> cardinal_sum (L (interval_co_0a n)f)
    in (iter a) +c (f a) = (iter (succ a)).
Proof.
move=> a f aB iter.
move: (partition_tack_on_intco aB) => pa.
set (g:= (L (interval_co_0a (succ a)) f)).
have fgg:fgraph g by rewrite /g; gprops.
have dg: domain g = (interval_co_0a (succ a)) by rewrite /g; bw.
rewrite - dg in pa.
rewrite /iter (cardinal_sum_assoc fgg pa); bw.
have adg: inc a (domain g) by rewrite dg; apply inc_a_interval_co_succ.
rewrite card_plus_pr0; bw; rewrite (trivial_cardinal_sum3 fgg adg).
rewrite interval_co_0a_restr // card_plus_pr2b //.
bw; rewrite {2} /g; bw; ue.
Qed.

Lemma induction_on_prod: forall a f, inc a Bnat ->
  let iter := fun n=> cardinal_prod (L (interval_co_0a n ) f)
    in (iter a) *c (f a) = (iter (succ a)).
Proof.
move=> a f aB iter.
move: (partition_tack_on_intco aB) => pa.
set (g:= (L (interval_co_0a (succ a)) f)).
have fgg:fgraph g by rewrite /g; gprops.
have dg: domain g = (interval_co_0a (succ a)) by rewrite /g; bw.
rewrite - dg in pa.
rewrite /iter (cardinal_prod_assoc fgg pa); bw.
have adg: inc a (domain g) by rewrite dg; apply inc_a_interval_co_succ.
rewrite card_mult_pr0; bw; rewrite trivial_cardinal_prod3 //.
rewrite interval_co_0a_restr // card_mult_pr2b //.
bw; rewrite {2} /g; bw; ue.
Qed.

Lemma fct_sum_rec0: forall f n, inc n Bnat ->
  cardinal_sum (L (interval_cc_0a n) f)
  = (cardinal_sum (L (interval_cc_1a n) f)) +c (f 0c).
Proof.
move=> f n nB.
move: (partition_tack_on_int nB) => pa.
set g:= L _ f.
have fgg:fgraph g by rewrite /g; gprops.
have dg: domain g = (interval_cc_0a n) by rewrite /g; bw.
rewrite - dg in pa.
rewrite (cardinal_sum_assoc fgg pa); bw.
have adg: inc 0c (domain g).
    rewrite dg ; srw; apply zero_smallest; fprops.
have Vz: (V 0c g = f 0c) by rewrite /g; bw; ue.
rewrite card_plus_pr0; bw.
rewrite (interval_int_restr _ nB) -/g -Vz.
by rewrite (trivial_cardinal_sum3 fgg adg) card_plus_pr2b.
Qed.

Lemma fct_sum_rec1: forall f n, inc n Bnat ->
  cardinal_sum (L (interval_Bnat 0c (succ n)) f)
  = (cardinal_sum (L (interval_Bnat 0c n) (fun i=> f (succ i)))) +c (f 0c).
Proof.
move=> f n nB.
have sB:inc (succ n) Bnat by hprops.
rewrite (fct_sum_rec0 _ sB); congr (_ +c _ ).
set i1:= interval_cc_1a _ ; set i0:= interval_Bnat _ _.
have p1: forall x : Set, inc x i0 -> inc (succ x) i1.
  move => t;rewrite /i0/i1 interval_Bnat_pr0 // interval_Bnat_pr1 //.
  move=> tn; split; [ apply succ_nonzero | rewrite -lt_n_succ_le1 //].
    by move: tn => [h_].
rewrite (cardinal_sum_commutative2 (I:=i0) (f:= succ)).
- congr (cardinal_sum _).
   apply L_exten1 =>//; move=> x xi0; bw; apply p1=>//.
- by move=> x xi; bw; apply p1.
- move=> u v ui vi susv; apply succ_injective1.
     move: (sub_interval_Bnat ui); fprops.
   move: (sub_interval_Bnat vi); fprops.
   done.
- move=> y; bw;rewrite /i1 interval_Bnat_pr1 //; move=> [nyz le_s].
  have fy:inc y Bnat by Bnat_tac.
  move: (predc_pr fy nyz)=> [pB ns].
  exists (predc y); split; last by [].
  rewrite /i0 interval_Bnat_pr0 // lt_n_succ_le1 //; fprops; ue.
- gprops.
Qed.

Lemma fct_sum_rev: forall f n, inc n Bnat ->
  let I := (interval_co_0a (succ n)) in
  cardinal_sum (L I f) = cardinal_sum (L I (fun i=> f (n -c i))).
Proof.
move=> f n nB I.
have snB: inc (succ n) Bnat by hprops.
set X := L I f.
pose g i := n -c i.
have p1: forall x, inc x I -> inc (n -c x) I.
  move=> x; rewrite /I; srw; move=> h; apply sub_le_symmetry =>//.
have->:(L I (fun i : Set => f (n -c i)) = L I (fun z => V (g z) X)).
  apply L_exten1 =>//; move=> x xI; rewrite /X /g; bw; apply p1 =>//.
apply cardinal_sum_commutative2.
- rewrite /g/X; bw.
- move=> x y; rewrite /I /g; srw; move=> xn yn h.
  by rewrite -(double_sub nB xn) -(double_sub nB yn) h.
- rewrite /X /g /I; bw; move=> y; srw => yn; rewrite -(double_sub nB yn).
  exists (n -c y); srw;ee; apply sub_le_symmetry =>//; Bnat_tac.
rewrite /X; gprops.
Qed.

EIII-5-4 Finite sequences


EIII-5-5 Characteristic functions on sets


Definition char_fun A B := BL (varianti A 1c 0c) B (doubleton 1c 0c).

Lemma char_fun_axioms: forall A B,
  transf_axioms (varianti A 1c 0c) B (doubleton 1c 0c).
Proof. move=> A B z zB; rewrite /varianti;Ytac1 zA; fprops. Qed.

Lemma char_fun_function: forall A B, is_function (char_fun A B).
Proof.
move=> A B; rewrite/char_fun; apply bl_function; apply char_fun_axioms.
Qed.

Lemma char_fun_W:forall A B x,
  inc x B -> W x (char_fun A B) = (varianti A 1c 0c x).
Proof. move=> A B x xB; rewrite /char_fun; aw; apply char_fun_axioms.
Qed.

Lemma char_fun_W_cardinal:forall A B x,
  inc x B -> is_cardinal (W x (char_fun A B)).
Proof.
move=> A B x xB; rewrite /char_fun; aw; last by apply char_fun_axioms.
rewrite /varianti;Ytac1 h; fprops.
Qed.

Lemma char_fun_W_a:forall A B x, sub A B -> inc x A ->
  W x (char_fun A B) = 1c.
Proof.
move=> A B x AB xA;move: (AB _ xA) => xB; rewrite char_fun_W //.
by apply varianti_in.
Qed.

Lemma char_fun_W_b:forall A B x, sub A B -> inc x (complement B A) ->
  W x (char_fun A B) = 0c.
Proof.
move=> A B x AB; srw; move=> [xB nxA]; rewrite char_fun_W //.
by apply varianti_out.
Qed.

Lemma char_fun_injective: forall A A' B, sub A B -> sub A' B ->
  (A=A') = (char_fun A B = char_fun A' B).
Proof.
move=> A A' B AB A'B; apply iff_eq; first by move => ->.
move=> scf; set_extens x => xs.
  ex_middle u.
  have xc: (inc x (complement B A')) by srw; split=>// ;apply AB.
  elim card_one_not_zero.
  by rewrite -(char_fun_W_a AB xs) -(char_fun_W_b A'B xc) scf.
ex_middle u.
have xc: (inc x (complement B A)) by srw; split=>// ;apply A'B.
elim card_one_not_zero.
by rewrite -(char_fun_W_a A'B xs) -(char_fun_W_b AB xc) scf.
Qed.

Lemma char_fun_W_aa:forall A x, inc x A ->
  W x (char_fun A A) = 1c.
Proof. by move=> A x xA ;rewrite char_fun_W // /varianti; Ytac0. Qed.

Lemma char_fun_W_bb:forall A x, inc x A ->
  W x (char_fun emptyset A) = 0c.
Proof.
move=> A x xA; rewrite char_fun_W // varianti_out //;case; case.
Qed.

Lemma char_fun_constant:forall A B, sub A B ->
  (forall x y, inc x B -> inc y B -> W x (char_fun A B) = W y (char_fun A B))
  -> (A=B \/ A = emptyset).
Proof.
move=> A B AB p.
case (emptyset_dichot A); first by auto.
move=> [u uA]; left.
apply extensionality=>// t tB.
ex_middle v.
have xc:inc t (complement B A) by srw; split =>//.
elim card_one_not_zero.
by rewrite - (char_fun_W_b AB xc)- (char_fun_W_a AB uA) (p _ _ (AB _ uA) tB).
Qed.

Lemma char_fun_complement: forall A B x, sub A B -> inc x B ->
  W x (char_fun (complement B A) B)
  = 1c -c (W x (char_fun A B)).
Proof.
move=> A B x AB xB.
have sc:sub (complement B A) B by apply sub_complement.
case (inc_or_not x A).
  move=> xA; rewrite char_fun_W_b //.
    rewrite char_fun_W_a //; symmetry; apply minus_n_nC; fprops.
  srw.
move=> h.
have xc: inc x (complement B A) by srw;auto.
rewrite char_fun_W_a // char_fun_W_b //.
symmetry;apply minus_n_0C; fprops.
Qed.

Lemma char_fun_inter: forall A A' B x, sub A B -> sub A' B -> inc x B ->
  W x (char_fun (intersection2 A A') B)
  = (W x (char_fun A B)) *c (W x (char_fun A' B)).
Proof.
move=> A A' B x AB A'B xB.
have Ha:sub (intersection2 A A') B.
  by move=> t; aw; move=> [tA _]; apply AB.
case (inc_or_not x A) => h.
  have cW:is_cardinal (W x (char_fun A' B)) by apply char_fun_W_cardinal.
  rewrite (char_fun_W_a AB h) one_unit_prodl //.
  case (inc_or_not x A') => xA.
    have aux:inc x (intersection2 A A') by apply intersection2_inc.
    rewrite char_fun_W_a // char_fun_W_a //.
  have aux:inc x (complement B (intersection2 A A')) by srw; aw; intuition.
  rewrite char_fun_W_b // char_fun_W_b //.
  srw; intuition.
have aux: inc x (complement B (intersection2 A A')) by srw; aw;intuition.
have aux2:inc x (complement B A) by srw; intuition.
by rewrite char_fun_W_b // char_fun_W_b // card_multC zero_prod_absorbing.
Qed.

Lemma char_fun_union: forall A A' B x, sub A B -> sub A' B -> inc x B ->
  (W x (char_fun (intersection2 A A') B))
  +c (W x (char_fun (union2 A A') B) )
  = (W x (char_fun A B)) +c (W x (char_fun A' B)).
Proof.
move=> A A' B x AB A'B xB.
have Ha:sub (intersection2 A A') B by move=> t; aw; intuition.
have Hb:sub (union2 A A') B by move=> t; rewrite union2_rw; case; intuition.
case (p_or_not_p (inc x A))=> xA.
  rewrite (char_fun_W_a AB xA).
  rewrite (char_fun_W_a Hb (union2_first A' xA)).
  case (p_or_not_p (inc x A'))=> xA'.
    rewrite (char_fun_W_a A'B xA') char_fun_W_a //.
     by apply intersection2_inc.
  rewrite card_plusC; apply f_equal.
  rewrite ! char_fun_W_b //; srw; ee; aw; intuition.
have Hc: inc x (complement B A) by srw; intuition.
have Hd:inc x (complement B (intersection2 A A')) by srw;aw; ee.
rewrite (char_fun_W_b AB Hc) (char_fun_W_b Ha Hd).
case (p_or_not_p (inc x A')) => aux.
  have xu: inc x (union2 A A') by apply union2_second.
  by rewrite (char_fun_W_a Hb xu) (char_fun_W_a A'B aux).
have xc:inc x (complement B (union2 A A')).
  by srw; ee; rewrite union2_rw; intuition.
have xba: inc x (complement B A') by srw.
by rewrite (char_fun_W_b Hb xc) (char_fun_W_b A'B xba).
Qed.

EIII-5-6 Euclidean Division


Lemma least_int_prop: forall prop:Set -> Prop,
  (forall x, prop x -> inc x Bnat) -> (exists x, prop x) ->
  prop 0c \/ (exists x, inc x Bnat & prop (succ x) & ~ prop x).
Proof.
move=> prop pi [x xp].
set (X:=Zo Bnat prop).
have sX: sub X Bnat by rewrite /X; apply Z_sub.
have nX: nonempty X by exists x; rewrite /X; apply Z_inc=>//; apply pi.
case (Bnat_wordered sX nX)=> q.
   by left; move: q; rewrite /X Z_rw; move=> [_].
right; move: q=> [a [aB []]]; rewrite /X Z_rw Z_rw; move => [_ psa] nh.
ex_tac.
Qed.

Lemma least_int_prop1: forall prop:Set -> Prop,
  (forall x, prop x -> inc x Bnat) -> ~(prop card_zero) ->
  (exists x, prop x) -> (exists x, inc x Bnat &prop(succ x) & ~ prop x).
Proof.
move=> p alp npz exp; case (least_int_prop alp exp)=>// h.
Qed.

Definition division_prop a b q r :=
  a = (b *c q) +c r & r <c b.

Definition card_rem0 a b:=
  choose (fun r => inc r Bnat & exists q, inc q Bnat & division_prop a b q r).
Definition card_quo0 a b:=
  choose (fun q => inc q Bnat & exists r, inc r Bnat & division_prop a b q r).
Definition card_rem a b := variant 0c a (card_rem0 a b) b.
Definition card_quo a b := variant 0c 0c (card_quo0 a b) b.

Definition BNdivides b a :=
  inc a Bnat & inc b Bnat & card_rem a b = 0c.

Notation "x %/c y" := (card_quo x y) (at level 50).
Notation "x %%c y" := (card_rem x y) (at level 50).
Notation "x %|c y" := (BNdivides x y) (at level 50).

Lemma division_prop_alt: forall a b q r, inc a Bnat-> inc b Bnat ->
  inc q Bnat-> inc r Bnat -> b <> 0c ->
  division_prop a b q r =
   ( (b *c q) <=c a & a <c (b *c succ q) & r = a -c (b *c q)).
Proof.
move=> a b q r aB bB qB rB nzb; rewrite /division_prop.
set (w:= b *c q).
rewrite Bsucc_rw // card_multC cardinal_distrib_prod_sum2 card_multC.
rewrite one_unit_prodl; fprops.
have wB: inc w Bnat by rewrite /w; hprops.
apply iff_eq.
  move=> [av lt];ee.
      rewrite av; apply sum_increasing3; fprops.
    rewrite av;apply finite_sum2_lt; fprops.
  by symmetry;apply card_sub_pr2 =>//; symmetry; rewrite card_plusC.
move=> [lwa [ltas rv]].
rewrite rv.
split; first by symmetry; apply card_sub_pr =>//.
move: ltas; rewrite -/w cardinal_lt_pr//; last by hprops.
move=> [c [cB [nzc s]]].
move: (card_sub_pr lwa); rewrite -rv=> aux.
rewrite cardinal_lt_pr; fprops ; exists c; split; [exact | split; [exact |]].
apply plus_simplifiable_left with w ; [exact | hprops| exact |].
by rewrite card_plusA aux.
Qed.

Lemma division_unique:forall a b q r q' r', inc a Bnat-> inc b Bnat ->
  inc q Bnat-> inc r Bnat -> inc q' Bnat-> inc r' Bnat -> b <> 0c ->
  division_prop a b q r -> division_prop a b q' r' ->
  (q = q' & r =r').
Proof.
move=> a b q r q' r' aB bB qB rB q'B r'B nbz.
rewrite division_prop_alt// division_prop_alt//.
move=> [le1 [lt1 e1]][le2 [lt2 e2]].
suff : q = q' by move=> h; split; ee; rewrite e1 e2 h.
move: (cardinal_le_lt_trans le1 lt2) => lt3.
move: (cardinal_le_lt_trans le2 lt1) => lt4.
move: (mult_lt_simplifiable bB q'B (BS_succ qB) nbz lt4).
move: (mult_lt_simplifiable bB qB (BS_succ q'B) nbz lt3).
srw; apply cardinal_antisymmetry1.
Qed.

Lemma division_exists:forall a b, inc a Bnat-> inc b Bnat ->
  b <> 0c -> exists q, exists r,
    (inc q Bnat & inc r Bnat & division_prop a b q r).
Proof.
move=> a b aB bB nzb.
set (pp:= fun q => inc q Bnat & a <c (b *c q)).
have np: (forall x, pp x -> inc x Bnat) by rewrite /pp; intuition.
have zp: (~ (pp 0c)).
   rewrite /pp; move => [_]; rewrite zero_prod_absorbing.
   move=> h;apply (zero_smallest1 h).
have ep: (exists x, pp x).
  rewrite /pp; exists (succ a); split; first by hprops.
  have aux: (succ a) <=c (b *c (succ a)).
     by rewrite card_multC; apply Bprod_increasing3; hprops.
  rewrite - lt_n_succ_le;hprops.
move: (least_int_prop1 np zp ep)=> [q [qB [q1 q2]]].
have p1: ((b *c q) <=c a).
  move: (Bnat_total_order (BS_mult bB qB) aB).
  rewrite /Bnat_lt /Bnat_le; case; first by move=> [_[_]].
  move=> [[_ [p1 p2]] p3]; elim q2; split; [exact | split ] =>//.
move: (card_sub_pr p1).
set (r:= a -c (b *c q)).
have rB: inc r Bnat by rewrite /r; apply (BS_sub aB (BS_mult bB qB)).
exists q; exists r; split; [exact | split ; [exact |]].
rewrite division_prop_alt; first (by move: q1=> []); exact.
Qed.

Lemma Bnat_division: forall a b, inc a Bnat-> inc b Bnat -> b <> 0c ->
  (inc (a %%c b) Bnat & (inc (a %/c b) Bnat) &
    (division_prop a b (a %/c b) (a %%c b))).
Proof.
move=> a b aB bB nbz.
have ->: (card_rem a b = card_rem0 a b) by rewrite /card_rem variant_if_not_rw.
have ->: (card_quo a b = card_quo0 a b) by rewrite /card_quo variant_if_not_rw.
move: (division_exists aB bB nbz) => [q [r [qB [rB p]]]].
have [rB' [q0 [q0B p1]]]: (inc (card_rem0 a b)Bnat & exists q,
    inc q Bnat &division_prop a b q (card_rem0 a b)).
  rewrite /card_rem0; apply choose_pr; exists r; ee; exists q; ee.
have [qB' [r0 [r0B p2]]]: (inc (card_quo0 a b) Bnat & exists r, inc r Bnat &
    division_prop a b (card_quo0 a b) r).
 rewrite /card_quo0; apply choose_pr; exists q; ee; exists r; ee.
move: (division_unique aB bB q0B rB' qB' r0B nbz p1 p2) => [aa bb].
rewrite bb; ee.
Qed.

Lemma card_quo_zero: forall a, a %/c 0c = 0c.
Proof. by move=> a; rewrite /card_quo/variant; Ytac0. Qed.

Lemma card_rem_zero: forall a, a %%c 0c = a.
Proof. by move=> a; rewrite /card_rem/variant; Ytac0. Qed.

Lemma BS_quo:forall a b, inc a Bnat-> inc b Bnat ->
  inc (a %/c b) Bnat.
Proof.
move => a b aB bB.
case (equal_or_not b 0c) => bnz.
  rewrite bnz card_quo_zero;fprops.
by move: (Bnat_division aB bB bnz) =>[_ [h _]].
Qed.

Lemma BS_rem:forall a b, inc a Bnat-> inc b Bnat ->
  inc (a %%c b) Bnat.
Proof.
move => a b aB bB.
case (equal_or_not b 0c) => bnz.
  rewrite bnz card_rem_zero //.
by move: (Bnat_division aB bB bnz) =>[h _].
Qed.

Hint Resolve BS_rem BS_quo: hprops.

Lemma Bnat_div_pr: forall a b, inc a Bnat-> inc b Bnat ->
  a = (b *c (a %/c b)) +c (a %%c b).
Proof.
move => a b aB bB.
case (equal_or_not b 0c) => bnz.
  rewrite bnz card_rem_zero card_quo_zero zero_prod_absorbing; aw; fprops.
by move: (Bnat_division aB bB bnz) =>[_ [_ [h _]]].
Qed.

Lemma Bnat_rem_pr: forall a b, inc a Bnat-> inc b Bnat -> b <> 0c ->
  (a %%c b) <c b.
Proof.
by move => a b aB bB bnz; move: (Bnat_division aB bB bnz) =>[_ [_ [_ h]]].
Qed.

Lemma Bnat_quorem_pr: forall a b q r,
  inc a Bnat-> inc b Bnat -> inc q Bnat -> inc r Bnat ->
  division_prop a b q r -> (q = a %/c b & r = a %%c b).
Proof.
move => a b q r aB bB qB rB p'.
case (equal_or_not b 0c) => bnz.
  move: p' => [_ ]; rewrite bnz =>h; case: (zero_smallest1 h).
move: (Bnat_division aB bB bnz)=> [rrB [qqB p]].
apply (division_unique aB bB qB rB qqB rrB bnz p' p).
Qed.

Lemma Bnat_quorem_pr0: forall a b q,
  inc a Bnat-> inc b Bnat -> inc q Bnat -> b <> 0c ->
   a = (b *c q) -> (q = a %/c b & 0c = a %%c b).
Proof.
move => a b q aB bB qB bnz p'.
apply Bnat_quorem_pr =>//.
  fprops.
red; rewrite -p' (Bnat0_unit_sumr aB).
split; [ trivial | split; [ apply zero_smallest; fprops | intuition ]].
Qed.

Lemma BNdivides_pr: forall a b,
  b %|c a -> a = b *c (a %/c b).
Proof.
move => a b [aB [bB dv]].
move: (Bnat_div_pr aB bB); rewrite dv.
by rewrite (Bnat0_unit_sumr (BS_mult bB (BS_quo aB bB))).
Qed.

Lemma BNdivides_pr1: forall a b, inc a Bnat -> inc b Bnat ->
  b %|c (b *c a).
Proof.
move=> a b aB bB.
move: (BS_mult bB aB) => pB.
red; ee.
case (equal_or_not b card_zero) => bnz.
  rewrite bnz card_multC zero_prod_absorbing card_rem_zero //.
move: (Bnat_quorem_pr0 pB bB aB bnz (refl_equal (card_mult b a))).
rewrite /BNdivides; intuition.
Qed.

Lemma BNdivides_pr2: forall a b q,
  inc a Bnat -> inc b Bnat -> inc q Bnat -> b <> 0c ->
  a = b *c q -> q = a %/c b.
Proof.
move => a b q aB bB qB nzb abq.
move: (Bnat_quorem_pr0 aB bB qB nzb abq).
intuition.
Qed.

Lemma BNdivides_one: forall a, inc a Bnat -> 1c %|c a.
Proof.
move=> a aB.
have -> :a = (card_mult card_one a) by rewrite one_unit_prodl; fprops.
apply (BNdivides_pr1 aB inc1_Bnat).
Qed.

Lemma BN_quo_one: forall a, inc a Bnat -> a %/c 1c = a.
Proof.
move=> a aB; symmetry; apply BNdivides_pr2; hprops.
apply card_one_not_zero.
rewrite one_unit_prodl; fprops.
Qed.

Lemma BNdivides_pr3: forall a b q,
   b %|c a -> q = a %/c b -> a = b *c q.
Proof. by move => a b q H aux; move: (BNdivides_pr H); rewrite -aux.
Qed.

Lemma BNdivides_pr4: forall b q, inc b Bnat -> inc q Bnat -> b <> 0c ->
   (b *c q) %/c b = q.
Proof. move=> *; symmetry; apply BNdivides_pr2 =>//; hprops.
Qed.

Lemma BNdivision_itself: forall a, inc a Bnat -> a <> 0c ->
  (a %|c a & a %/c a = 1c).
Proof.
move=> a aB anz; rewrite /BNdivides.
have aux:a = a *c 1c by rewrite one_unit_prodr; fprops.
move: (Bnat_quorem_pr0 aB aB inc1_Bnat anz aux).
by move=> [p1 p2]; ee.
Qed.

Lemma BNdivides_itself: forall a,inc a Bnat -> a <> card_zero ->
   a %|c a.
Proof. by move=> a aB anz; move: (BNdivision_itself aB anz) => [h _].
Qed.

Lemma BNquo_itself: forall a, inc a Bnat -> a <> 0c ->
   a %/c a = 1c.
Proof. by move=> a aB anz; move: (BNdivision_itself aB anz) => [ _ h].
Qed.

Lemma BNdivision_of_zero: forall a, inc a Bnat ->
  (a %|c 0c & 0c %/c a = 0c).
Proof.
move=> a aB; rewrite /BNdivides.
have aux: 0c = a *c 0c by rewrite zero_prod_absorbing.
case (equal_or_not a 0c) => anz.
  rewrite anz card_quo_zero card_rem_zero; ee.
move: (Bnat_quorem_pr0 inc0_Bnat aB inc0_Bnat anz aux).
by move=> [p1 p2]; ee.
Qed.

Lemma BNdivides_trans: forall a b a',
  a %|c a'-> b %|c a -> b %|c a'.
Proof.
move=> a b a' d1 d2.
rewrite (BNdivides_pr d1) {1} (BNdivides_pr d2) - card_multA.
move: d1 d2 => [p1 [p2 _]] [p3 [p4 _]].
apply BNdivides_pr1 =>//; hprops.
Qed.

Lemma BNdivides_trans1: forall a b a',
  a %|c a' -> b %|c a
  -> a' %/c b = (a' %/c a) *c (a %/c b).
Proof.
move=> a b a' d1 d2.
move: (BNdivides_pr d1).
rewrite {1} (BNdivides_pr d2) - card_multA => h.
case (equal_or_not b 0c) => bnz.
 by rewrite bnz card_quo_zero card_quo_zero zero_prod_absorbing.
move: d1 d2 => [a'B [aB _]][_ [bB _]].
have pB: inc ((a %/c b) *c (a' %/c a)) Bnat by hprops.
rewrite - (BNdivides_pr2 a'B bB pB bnz h).
by rewrite card_multC.
Qed.

Lemma BNdivides_trans2: forall a b c, inc c Bnat ->
  b %|c a-> b %|c (a *c c).
Proof.
move=> a b c cB ba.
have aB: (inc a Bnat) by move: ba => [h _].
move: (BNdivides_pr1 cB aB) => aux.
apply (BNdivides_trans aux ba).
Qed.

Lemma card_quo_simplify: forall a b c,
  inc a Bnat -> inc b Bnat -> inc c Bnat -> b <> 0c -> c <> 0c ->
  (a *c c) %/c (b *c c) = a %/c b.
Proof.
move=> a b c aB bB cB bnz cnz.
move: (Bnat_division aB bB bnz)=> [rB [qB []]].
set q:= (a %/c b); set r := a %%c b.
move=> e1 lrb; symmetry.
move: (BS_mult aB cB)(BS_mult bB cB)(BS_mult (BS_rem aB bB) cB) => p1 p2 p3.
have dv: (division_prop (a *c c) (b *c c) q (r *c c)).
  split.
     rewrite (card_multC b c) - card_multA (card_multC c _).
    by rewrite - cardinal_distrib_prod_sum2 e1.
  rewrite (card_multC b c) card_multC.
  apply finite_prod2_lt =>//; fprops.
by move: (Bnat_quorem_pr p1 p2 qB p3 dv) => [res _].
Qed.

Lemma divides_and_sum: forall a a' b, b %|c a -> b %|c a'
  -> (b %|c (a +c a') &
    (a +c a') %/c b = (a %/c b) +c (a' %/c b)).
Proof.
move=> a a' b d1 d2.
move: (BNdivides_pr d1)(BNdivides_pr d2)=> am bm.
set s := (a %/c b) +c (a' %/c b).
have ->: (a +c a' = b *c s).
  by rewrite /s cardinal_distrib_prod_sum3 -am -bm.
move: d1 d2 => [aB [bB _]][a'B _].
have sB:inc s Bnat by rewrite /s; hprops.
split; first by apply BNdivides_pr1 =>//.
case (equal_or_not b 0c) => bs.
  rewrite {2} /s bs ! card_quo_zero; aw; fprops.
symmetry; apply (BNdivides_pr2 (BS_mult bB sB) bB sB bs).
apply refl_equal.
Qed.

Lemma distrib_prod2_sub: forall a b c, inc a Bnat -> inc b Bnat -> inc c Bnat
  -> c <=c b ->
  a *c (b -c c) = (a *c b) -c (a *c c).
Proof.
move=> a b c aB bB cB le; symmetry; apply card_sub_pr2; hprops.
rewrite - cardinal_distrib_prod_sum3 card_plusC card_sub_pr //.
Qed.

Lemma divides_and_difference: forall a a' b,
  a' <=c a -> b %|c a -> b %|c a'
  -> (b %|c (a -c a') & (a' %/c b) <=c (a %/c b) &
    (a -c a') %/c b = (a %/c b) -c (a' %/c b)).
Proof.
move=> a a' b le d1 d2.
move: (BNdivides_pr d1)(BNdivides_pr d2)=> am bm.
set s := (a %/c b) -c (a' %/c b).
move: d1 d2 => [aB [bB _]][a'B _].
move: (BS_quo aB bB)(BS_quo a'B bB) => q1 q2.
have q3: (a' %/c b) <=c (a %/c b).
  case (equal_or_not b card_zero) => bnz.
    rewrite bnz ! card_quo_zero; fprops.
  rewrite am bm in le.
  apply (mult_le_simplifiable bB q2 q1 bnz le).
have ->: (a -c a' = b *c s).
  by rewrite /s distrib_prod2_sub // -am -bm.
have sB:inc s Bnat by rewrite /s; hprops.
ee; first by apply BNdivides_pr1 =>//.
case (equal_or_not b card_zero) => bs.
  rewrite {2} /s bs ! card_quo_zero minus_n_nC; fprops.
by symmetry; apply (BNdivides_pr2 (BS_mult bB sB) bB sB bs); apply refl_equal.
Qed.

EIII-5-7 Expansion to base b


Lemma b_power_k_large: forall a b, inc a Bnat -> inc b Bnat ->
  1c <c b -> a <> 0c -> exists k,
    inc k Bnat & (b ^c k) <=c a & a <c (b ^c (succ k)).
Proof.
move=> a b aB bB l1b naz.
set (prop:= fun k=> inc k Bnat & a <c (b ^c k)).
have pB:(forall x, prop x -> inc x Bnat) by rewrite /prop=> k; ee.
have pz: ~(prop card_zero).
  rewrite /prop power_x_0; move=> [_ a1].
   move: (one_small_cardinal (Bnat_is_cardinal aB) naz)=> h; co_tac.
have exp:(exists x, prop x).
  by exists a; red; ee; apply lt_a_power_b_a.
move: (least_int_prop1 pB pz exp)=> [k [kN [pks npk]]].
ex_tac;split; last by move: pks => [_ res].
case (Bnat_total_order (BS_pow bB kN) aB).
  rewrite /Bnat_le; intuition.
move=>h; elim npk; split; first by exact.
move: h; rewrite /Bnat_lt /Bnat_le /cardinal_lt.
by move => [[_ [_ p1]] p2]; split.
Qed.

Definition is_expansion f b k :=
  inc b Bnat & inc k Bnat & 1c <c b &
  fgraph f & domain f = interval_co_0a k &
  forall i, inc i (domain f) -> (V i f) <c b.

Definition expansion_value f b :=
  cardinal_sum (L (domain f) (fun i=> (V i f) *c (b ^c i))).

Section Base_b_expansion.
Variables f g b k k': Set.
Variable Exp: is_expansion f b k.
Variable Expg: is_expansion g b (succ k').

Lemma is_expansion_prop0: forall i,
   (inc i (domain f)) = i <c k.
Proof. move: Exp => [_ [kn [_ [_ [d _]]]]] i; rewrite d; srw.
Qed.

Lemma is_expansion_prop1: forall i,
  i <c k -> inc (V i f) Bnat.
Proof.
move: Exp => [bN [kn [_ [_ [d v]]]]] i.
rewrite -is_expansion_prop0 //; move=> idf.
move: (v _ idf) => [h _ ]; Bnat_tac.
Qed.

Lemma is_expansion_prop2:
   finite_int_fam (L (domain f) (fun i=> (V i f) *c (b ^c i))).
Proof.
move: Exp => [bB [kB [b2 [fgf [df vf]]]]].
red;bw;ee.
  move=> i idf; bw.
  rewrite is_expansion_prop0 in idf.
  rewrite - inc_Bnat; apply BS_mult.
  apply (is_expansion_prop1) =>//.
  have iB: (inc i Bnat) by move: idf => [idf _]; Bnat_tac.
  hprops.
rewrite df; red; rewrite cardinal_interval_co_0a1 -? inc_Bnat //.
Qed.

Lemma is_expansion_prop3:
  inc (expansion_value f b) Bnat.
Proof.
rewrite /expansion_value; apply finite_sum_finite.
apply is_expansion_prop2.
Qed.

Lemma is_expansion_prop4: is_cardinal k' -> inc k' Bnat.
Proof.
move: Expg => [_ [kB' _ ]] ck; apply BS_nsucc => //.
Qed.

Lemma is_expansion_prop5: is_cardinal k' ->
  is_expansion (restr g (interval_co_0a k')) b k'.
Proof.
move => bk'; move: (is_expansion_prop4 bk')=> k'n.
move: (interval_co_0a_increasing k'n) => subi.
move: Expg => [bB [kB [b2 [fgf [df vf]]]]].
have dr: (domain (restr g (interval_co_0a k')) = interval_co_0a k').
  rewrite restr_domain1 //; ue.
red; ee.
rewrite dr; move=> i ic; bw; try ue.
by apply vf; rewrite df; apply subi.
Qed.

Lemma is_expansion_prop6: is_cardinal k' -> inc (V k' g) Bnat.
Proof.
move => bk'; move: (is_expansion_prop4 bk')=> k'n.
move: Expg => [bN [kn [_ [_ [d v]]]]].
have kd: (inc k' (domain g)) by rewrite d; srw; fprops.
move: (v _ kd)=> [h _]; Bnat_tac.
Qed.

Lemma is_expansion_prop7: is_cardinal k' ->
  (expansion_value g b) =
   (expansion_value (restr g (interval_co_0a k')) b)
   +c (V k' g *c (b ^c k')).
Proof.
move => bk'; move: (is_expansion_prop4 bk')=> k'n.
pose h i := (V i g) *c (b ^c i).
have ->: (V k' g) *c (b ^c k') = h k' by [].
rewrite /expansion_value.
move: (induction_on_sum h k'n); simpl.
have dg: domain g = (interval_co_0a (succ k'))
  by move: Expg=> [_[_[_[_ [dg _]]]]].
rewrite -dg; set kk:= (L (domain g) h) ; move=> <-.
set f1:= L _ _; set f2:= L _ _.
suff: f1 = f2 by move => ->.
move: (interval_co_0a_increasing k'n) => ii.
have fgg: fgraph g by move : Expg=> []; intuition.
apply L_exten1.
  rewrite restr_domain1 // dg //.
move=> x xdf; bw; ue.
Qed.
End Base_b_expansion.

Lemma is_expansion_prop8: forall f b k x,
  let h:= L (interval_co_0a (succ k)) (fun i=> Yo (i=k) x (V i f)) in
 is_expansion f b k ->
   inc x Bnat -> x <c b ->
    (is_expansion h b (succ k) &
      expansion_value h b =
      (expansion_value f b) +c ((b ^c k) *c x)).
Proof.
move=> f b k x h Expf xB xb.
move: Expf => [bB [kB [b2 [fgf [df vf]]]]].
have eg: (is_expansion h b (succ k)).
  red; rewrite /h; bw; rewrite - !inc_Bnat; ee.
  move=> i idh; bw; Ytac ik => //; apply vf.
  rewrite df; move: (interval_co_pr4 kB)=> [p1 p2].
  rewrite - p1 in idh; awi idh; case idh; intuition.
have ck: is_cardinal k by fprops.
rewrite (is_expansion_prop7 eg ck).
split; first by exact.
have ->:(restr h (interval_co_0a k)= f).
  rewrite /h; symmetry.
  move: (interval_co_0a_increasing kB) => si.
  apply fgraph_exten=> //.
     apply restr_fgraph; gprops.
    rewrite restr_domain1 //; gprops; bw.
  rewrite df; move=> y ydf; bw; gprops.
  by Ytac yk; move: ydf; srw;move => [_]; rewrite yk.
have ->: (V k h *c (b ^c k) = (b ^c k) *c x).
  rewrite card_multC /h L_V_rw; first by Ytac0.
  apply inc_a_interval_co_succ =>//.
done.
Qed.

Lemma is_expansion_prop9: forall f b k, is_expansion f b k ->
   (expansion_value f b) <c (b ^c k).
Proof.
move=> f b k Exp.
have kB: inc k Bnat by move: Exp => [bB [kB _]].
move: k kB f Exp.
apply: cardinal_c_induction.
  rewrite power_x_0.
  move=> g [bB [_ [b2 [fgf [df _]]]]].
  rewrite /expansion_value df emptyset_interval_00 trivial_cardinal_sum.
    apply zero_lt_one.
  bw.
move=> n nB pn g eg.
have cn: (is_cardinal n) by fprops.
rewrite (is_expansion_prop7 eg cn).
move: (is_expansion_prop5 eg cn) => er.
move: (pn _ er).
move : eg => [bB [_ [b2 [fgf [df vg]]]]].
rewrite pow_succ //.
move: (is_expansion_prop3 er).
set (a0:= expansion_value (restr g (interval_co_0a n)) b).
set (b0:= b ^c n).
have p1: ((V n g) <c b).
  by apply vg; rewrite df; apply inc_a_interval_co_succ.
have cB: inc (V n g) Bnat by move: p1 => [aux _]; Bnat_tac.
have b0B: inc b0 Bnat by rewrite /b0; apply BS_pow.
set c0:= (V n g) *c b0.
move => a0B h.
have le1: (a0 +c c0) <c (b0 +c c0).
  apply finite_sum3_lt; rewrite /c0; hprops.
suff le2: (b0 +c c0) <=c (b0 *c b) by co_tac.
have cb: is_cardinal b0 by fprops.
have ->: b0 +c c0 = b0 *c (succ (V n g)).
  rewrite (Bsucc_rw cB) cardinal_distrib_prod_sum3 one_unit_prodr.
    rewrite card_multC card_plusC //.
    exact.
apply product_increasing2; fprops.
rewrite -lt_n_succ_le in p1; fprops.
Qed.

Lemma is_expansion_prop10: forall f b k, is_cardinal k ->
  is_expansion f b (succ k) ->
  division_prop (expansion_value f b) (b ^c k) (V k f)
  (expansion_value (restr f (interval_co_0a k)) b).
Proof.
move=> f b k ck ie.
split.
  rewrite card_plusC card_multC.
  apply is_expansion_prop7 =>//.
apply is_expansion_prop9; apply is_expansion_prop5 =>//.
Qed.

Lemma is_expansion_unique: forall f g b k,
  is_expansion f b k -> is_expansion g b k ->
  expansion_value f b = expansion_value g b -> f = g.
Proof.
move=> f g b k ef eg sm.
have kB: inc k Bnat by move: eg => [_ [kB _ ]].
move: k kB f g sm ef eg.
apply: (cardinal_c_induction).
  rewrite /is_expansion emptyset_interval_00.
  move=> f1 g1 sv [bB [_ [_ [fgf [df _]]]]] [_ [_ [_ [fgg [dg _]]]]].
  apply fgraph_exten; [exact | exact | ue | ].
  rewrite df;move=> x xdf; elim (emptyset_pr xdf).
move=> n nB pn f g sv ef eg.
have cn: is_cardinal n by fprops.
move: (is_expansion_prop10 cn ef) => p1.
move: (is_expansion_prop10 cn eg) => p2.
rewrite sv in p1.
move: (is_expansion_prop5 ef cn) => er1.
move: (is_expansion_prop5 eg cn) => er2.
move:(is_expansion_prop3 eg) => egB.
have bB: inc b Bnat by move: ef => [bB _].
have cpB: inc (b ^c n) Bnat by hprops.
move: (is_expansion_prop6 ef cn) => q1B.
move: (is_expansion_prop6 eg cn) => q2B.
move: (is_expansion_prop3 er1) => r1B.
move: (is_expansion_prop3 er2) => r2B.
have bnz: (b ^c n <> card_zero).
  have b0: (0c <c b).
    have aux: 0c <=c 1c by fprops.
    move: ef => [_ [_ [n1b _]]]; co_tac.
  apply non_zero_apowb;fprops.
  move: b0 => [_]; intuition.
move: (division_unique egB cpB q1B r1B q2B r2B bnz p1 p2)=> [pt r].
have aux: (restr f (interval_co_0a n) = restr g (interval_co_0a n)).
  apply pn =>//.
move:ef eg=> [_ [_ [_ [fgf [df _]]]]] [_ [_ [_ [fgg [dg _]]]]].
apply fgraph_exten; [ exact | exact | ue | ].
move=> x xdf.
case (equal_or_not x n); first by move => ->.
move=> xn.
have xc: (inc x (interval_co_0a n)).
  move: xdf; rewrite df (interval_co_0a_pr2 _ nB)
      (interval_co_0a_pr2 _ (BS_succ nB)) (lt_is_le_succ _ nB)=> lxn; red;ee.
have <-: (V x (restr f (interval_co_0a n)) = V x f).
  bw; rewrite df; apply interval_co_0a_increasing =>//.
have <-: (V x (restr g (interval_co_0a n)) = V x g).
  bw; rewrite dg; apply interval_co_0a_increasing =>//.
ue.
Qed.

Lemma is_expansion_prop11: forall f g b k, is_cardinal k ->
  is_expansion f b (succ k) -> is_expansion g b (succ k) ->
  (V k f) <c (V k g) ->
  (expansion_value f b) <c (expansion_value g b).
Proof.
move=> f g b k ck ef eg ltk.
rewrite (is_expansion_prop7 ef ck) (is_expansion_prop7 eg ck).
move: (is_expansion_prop5 ef ck) => ef1.
move: (is_expansion_prop5 eg ck) => eg1.
move: (is_expansion_prop9 ef1).
move: (is_expansion_prop3 ef1).
move: (is_expansion_prop3 eg1).
set u:= expansion_value _ _; set v:= expansion_value _ _.
move=> uB vB vp.
move: (is_expansion_prop4 ef ck) => kB.
move: (is_expansion_prop6 ef ck) => fB.
move: (is_expansion_prop6 eg ck) => gB.
set (B:= card_pow b k).
have bB: inc b Bnat by move: ef => [h _].
have BB: (inc B Bnat) by rewrite /B; hprops.
have bb: (B <=c B) by rewrite /B;fprops.
apply cardinal_lt_le_trans with (B +c ((V k f) *c B)).
  apply finite_sum3_lt; hprops.
apply cardinal_le_transitive with ((V k g) *c B); last first.
  rewrite card_plusC; apply sum_increasing3; fprops.
rewrite - (lt_n_succ_le fB gB) in ltk.
move: (product_increasing2 ltk bb).
rewrite (Bsucc_rw fB) cardinal_distrib_prod_sum2.
by rewrite (one_unit_prodl (Bnat_is_cardinal BB)) card_plusC.
Qed.

Lemma is_expansion_restr1: forall f b k l,
  is_expansion f b k -> l <=c k ->
  is_expansion (restr f (interval_co_0a l)) b l.
Proof.
move=> f b k l [bB [bK [b2 [fgf [df vg]]]]] lk.
have lB: inc l Bnat by Bnat_tac.
have slk: sub (interval_co_0a l) (interval_co_0a k).
  move=> t; srw => tl; co_tac.
have dr: domain (restr f (interval_co_0a l)) = interval_co_0a l.
  by rewrite restr_domain1 //; ue.
red; ee.
rewrite dr;move=> i idf; bw; [apply vg|]; ue.
Qed.

Lemma is_expansion_restr2: forall f b k l,
  is_expansion f b k -> l <=c k ->
  (forall i, l <=c i -> i <c k -> V i f = 0c) ->
  expansion_value (restr f (interval_co_0a l)) b = expansion_value f b.
Proof.
move=> f b k l ef lk p.
set g := fun i => expansion_value (restr f (interval_co_0a i)) b.
have <-: (g k = expansion_value f b).
  move: ef => [bB [bK [b2 [fgf [df vg]]]]].
  rewrite /g -df restr_to_domain //.
set r := fun i=> g l = g i.
change (r k).
have kB: inc k Bnat by move: ef => [bB [bK _]].
have lB: inc l Bnat by Bnat_tac.
have: inc k (interval_cc Bnat_order l k).
   rewrite Bnat_interval_cc_pr1; ee.
apply (cardinal_c_induction3_v (r:=r) lB kB); first by trivial.
move=> i; rewrite (Bnat_interval_co_pr1 _ lB kB).
rewrite /r /g; move=> [li ik]; move=> ->.
have iB: inc i Bnat by move: ik => [ik _]; Bnat_tac.
have ci: is_cardinal i by fprops.
have ik': (succ i) <=c k by rewrite lt_n_succ_le0 //; fprops.
move:(is_expansion_restr1 ef ik').
set F := (restr f (interval_co_0a (succ i))).
move=> eF.
rewrite (is_expansion_prop7 eF ci).
set G := (restr F (interval_co_0a i)).
have si: sub (interval_co_0a (succ i)) (domain f).
  move: ef => [bB [bK [b2 [fgf [df vg]]]]].
  rewrite df; move=> t; srw; hprops; move=> ti; co_tac.
have ->: (restr f (interval_co_0a i) = G).
 move: ef => [bB [bK [b2 [fgf [df vg]]]]].
 rewrite /G /F double_restr //.
 apply interval_co_0a_increasing =>//.
have ->: (V i F = card_zero).
  move: ef => [bB [bK [b2 [fgf [df vg]]]]].
  rewrite /F; bw; try apply p=>//.
  apply inc_a_interval_co_succ =>//.
rewrite card_multC zero_prod_absorbing zero_unit_sumr //.
rewrite /expansion_value/cardinal_sum; fprops.
Qed.

Lemma is_expansion_prop12: forall f g b kf kg l n,
  n <=c kf -> n <=c kg -> l <c n ->
  (forall i, n <=c i -> i <c kf -> V i f = 0c) ->
  (forall i, n <=c i -> i <c kg -> V i g = 0c) ->
  (forall i, l <c i -> i <c n -> V i f = V i g) ->
  is_expansion f b kf -> is_expansion g b kg ->
  (V l f) <c (V l g) ->
  (expansion_value f b) <c (expansion_value g b).
Proof.
move=> f g b kf kg l n nkf nkg ln pf pg pfg ef eg vl.
move: (is_expansion_restr1 ef nkf).
rewrite -(is_expansion_restr2 ef nkf pf).
move: (is_expansion_restr1 eg nkg).
rewrite -(is_expansion_restr2 eg nkg pg).
set F := (restr f (interval_co_0a n)).
set G := (restr g (interval_co_0a n)).
move=> eG eF; clear pf pg.
have kfB: inc kf Bnat by move: ef => [_ [h _]].
have nB: inc n Bnat by Bnat_tac.
have lB: inc l Bnat by move: ln => [ln _]; Bnat_tac.
have pFG: forall i, l <c i -> i <c n -> V i F = V i G.
  move=> i i1 i2; move: (pfg _ i1 i2).
  have ii: inc i (interval_co_0a n) by srw.
  move: ef eg => [_ [_ [_ [fgf [df vf]]]]] [_ [kgB [b2 [fgg [dg vg]]]]].
  rewrite /F/G; bw.
   rewrite dg; move=> t; srw =>//;move=> gh; co_tac.
   rewrite df; move=> t; srw =>//; move=> fh; co_tac.
have vL: (V l F) <c (V l G).
  have ii: inc l (interval_co_0a n) by srw.
  move: ef eg => [_ [_ [_ [fgf [df vf]]]]] [_ [kgB [b2 [fgg [dg vg]]]]].
  rewrite /F/G; bw.
   rewrite dg; move=> t; srw =>//;move=> gh; co_tac.
   rewrite df; move=> t; srw =>//; move=> fh; co_tac.
clear pfg vl ef eg.
set fi := fun i => expansion_value (restr F (interval_co_0a i)) b.
set gi := fun i => expansion_value (restr G (interval_co_0a i)) b.
have <-: (fi n = expansion_value F b).
  move: eF => [bB [bK [b2 [fgf [df vg]]]]].
  rewrite /fi -df restr_to_domain //.
have <-: (gi n = expansion_value G b).
  move: eG => [bB [bK [b2 [fgf [df vg]]]]].
  rewrite /gi -df restr_to_domain //.
pose r i := cardinal_lt (fi i) (gi i).
change (r n).
have cl2: inc (succ l) Bnat by hprops.
have cl0: (succ l) <=c n by rewrite lt_n_succ_le0; fprops.
have rz: (r (succ l)).
  rewrite /r /fi /gi.
  have cl: is_cardinal l by fprops.
  have cl1: finite_c l by rewrite -inc_Bnat.
  move: (is_expansion_restr1 eF cl0) => ef1.
  move: (is_expansion_restr1 eG cl0) => eg1.
  apply (is_expansion_prop11 cl ef1 eg1).
  move: eF eG => [_ [_ [_ [fgf [df vf]]]]] [_ [kgB [b2 [fgg [dg vg]]]]].
  have il: inc l (interval_co_0a (succ l)).
     apply inc_a_interval_co_succ =>//.
  bw.
  rewrite dg; move=> t; srw =>//; move=> gh; co_tac.
  rewrite df; move=> t; srw =>//; move=> gh; co_tac.
have: inc n (interval_cc Bnat_order (succ l) n).
   rewrite Bnat_interval_cc_pr1; ee.
apply (cardinal_c_induction3_v (r:=r) cl2 nB rz).
clear rz.
move=> i; rewrite (Bnat_interval_co_pr1 _ cl2 nB).
rewrite /r /fi /gi; move=> [li ik] ri.

have iB: inc i Bnat by move: ik => [ik _]; Bnat_tac.
have ci: is_cardinal i by fprops.
have ik': (succ i) <=c n by rewrite lt_n_succ_le0 //; fprops.
move:(is_expansion_restr1 eF ik').
move:(is_expansion_restr1 eG ik').
set f1 := (restr F (interval_co_0a (succ i))).
set g1 := (restr G (interval_co_0a (succ i))).
move=> ef1 eg1; move: ri.
rewrite (is_expansion_prop7 ef1 ci).
rewrite (is_expansion_prop7 eg1 ci).
set f2 := (restr f1 (interval_co_0a i)).
set g2 := (restr g1 (interval_co_0a i)).
have si: sub (interval_co_0a (succ i)) (domain F).
  move: eF => [bB [bK [b2 [fgf [df vg]]]]].
  rewrite df; move=> t; srw; hprops; move=> ti; co_tac.
have sj: sub (interval_co_0a (succ i)) (domain G).
  move: eG => [bB [bK [b2 [fgf [df vg]]]]].
  rewrite df; move=> t; srw; hprops; move=> ti; co_tac.
have ->: (restr F (interval_co_0a i) = f2).
 move: eF => [bB [bK [b2 [fgf [df vg]]]]].
 rewrite /f2 /f1 double_restr //.
 apply interval_co_0a_increasing =>//.
have ->: (restr G (interval_co_0a i) = g2).
 move: eG => [bB [bK [b2 [fgf [df vg]]]]].
 rewrite /g2 /g1 double_restr //.
 apply interval_co_0a_increasing =>//.
have ->: (V i f1 = V i g1).
  have ii:inc i (interval_co_0a (succ i)) by apply inc_a_interval_co_succ =>//.
  move: eF => [bB [bK [b2 [fgf [df vg]]]]].
  move: eG => [_ [_ [_ [fgf' [df' vg']]]]].
  rewrite /f1 /g1; bw; try apply pFG=>//.
  rewrite -lt_n_succ_le0 //; fprops.
have isi: i <=c (succ i) by apply is_le_succ.
move:(is_expansion_restr1 eg1 isi) => ef2.
move:(is_expansion_restr1 ef1 isi) => eg2.
apply finite_sum3_lt.
    apply (is_expansion_prop3 ef2).
  have bB: inc b Bnat by move: ef1 => [ h _].
  apply BS_mult; hprops.
  apply (is_expansion_prop1 ef1); fprops.
  apply is_lt_succ; fprops.
apply (is_expansion_prop3 eg2).
Qed.

Lemma is_expansion_prop13: forall f g b kf kg l,
  kf <=c l -> l <c kg ->
  is_expansion f b kf -> is_expansion g b kg ->
  V l g <> 0c ->
  (expansion_value f b) <c (expansion_value g b).
Proof.
move=> f g b kf kg l le1 le2 ef eg vnz.
apply cardinal_lt_le_trans with (b ^c kf).
apply: (is_expansion_prop9 ef).
move: eg => [bB [bK [b2 [fgf [df vg]]]]].
rewrite /expansion_value.
set F:= L _ _.
have fgF: fgraph F by rewrite /F;gprops.
have dF: domain F = interval_co_0a kg by rewrite /F -df; bw.
have alc: (forall x, inc x (domain F) -> is_cardinal (V x F)).
  rewrite dF -df /F => x xdf; bw; fprops.
have ldf: inc l (domain F) by rewrite dF; srw.
set j:= singleton l.
have sj: sub j (domain F) by rewrite /j dF; move=> t; aw; move=> ->; ue.
move: (sum_increasing1 fgF alc sj).
rewrite (trivial_cardinal_sum3 fgF ldf) (cardinal_of_cardinal (alc _ ldf)).
suff: (b ^c kf) <=c (V l F) by move => p1 p2; co_tac.
have ldg: inc l (domain g) by rewrite df; srw.
rewrite /F; bw.
have bnz: b<> card_zero by move=>h; rewrite h in b2; elim (zero_smallest1 b2).
apply cardinal_le_transitive with (b ^c l).
  apply power_increasing1; fprops.
rewrite card_multC.
apply product_increasing3 => //.
  rewrite /card_pow; fprops.
by move: (vg _ ldg)=> [[h _] _].
Qed.

Lemma is_expansion_prop14: forall f g b kf kg,
  is_expansion f b kf -> is_expansion g b kg ->
  (expansion_value f b) <c (expansion_value g b) ->
  (exists l, kf <=c l & l <c kg & V l g <> 0c)
 \/ (
  exists l, exists n,
  (n <=c kf & n <=c kg & l <c n &
  (forall i, n <=c i -> i <c kf -> V i f = 0c) &
  (forall i, n <=c i -> i <c kg -> V i g = 0c) &
  (forall i, l <c i -> i <c n -> V i f = V i g) &
   (V l f)<c (V l g))).
Proof.
move=> f g b kf kg ef eg lt.
have kfb:inc kf Bnat by move: ef => [_ [h _]].
have kgb:inc kg Bnat by move: eg => [_ [h _]].
have [n [nf [ng nfg]]]: exists n,
   n <=c kf & n <=c kg & (n = kf \/ n = kg).
  have k1:is_cardinal kf by fprops.
  have k2:is_cardinal kg by fprops.
  case (cardinal_le_total_order3 k1 k2)=> aux; [exists kf | exists kg];ee.
case (p_or_not_p (forall i, n <=c i -> i <c kg -> V i g = 0c)); last first.
  move=> h.
  have [i [ni [ik Vz]]]: (exists i, n <=c i & i <c kg & V i g <> 0c).
    apply exists_proof; move=> h'.
    elim h; move=> i i1 i2.
    case (p_or_not_p (V i g = 0c)) =>// h''; elim (h' i); ee.
  left; exists i; ee.
  case nfg; [by move=> <- | move=> nk; rewrite nk in ni; co_tac ].
move=> pB; right.
case (p_or_not_p (forall i, n <=c i -> i <c kf -> V i f = 0c)); last first.
  move=> h.
  have [i [ni [ik Vz]]]: (exists i, n <=c i & i <c kf & V i f <> 0c).
    apply exists_proof; move=> h'; elim h; move=> i i1 i2.
    case (p_or_not_p (V i f = 0c)) =>// h''; elim (h' i); ee.
  have fi: kg <=c i.
  case nfg; [move=> nk; rewrite nk in ni; co_tac | by move=> <-].
  move: (is_expansion_prop13 fi ik eg ef Vz) =>[ br _].
  co_tac.
move=> pA.
have nB: inc n Bnat by case nfg; move=> -> //.
have pC: exists l, l <c n & (V l f) <> (V l g).
  move: (is_expansion_restr2 ef nf pA) => hf.
  move: (is_expansion_restr2 eg ng pB) => hg.
  apply exists_proof.
  move=> h.
  have eq: (restr f (interval_co_0a n) = restr g (interval_co_0a n)).
    move: ef eg => [bB [bK [b2 [fgf [df vf]]]]][_ [_ [_ [fgg [dg vg]]]]].
    have s1: sub (interval_co_0a n) (domain g).
      rewrite dg; move=> t; srw; move=> h1; co_tac.
    have s2: sub (interval_co_0a n) (domain f).
      rewrite df; move=> t; srw; move=> h1; co_tac.
    have drf: (domain (restr f (interval_co_0a n)) = interval_co_0a n).
      rewrite restr_domain1 //; fprops.
    have drg: (domain (restr g (interval_co_0a n)) = interval_co_0a n).
      rewrite restr_domain1 //; fprops.
    apply fgraph_exten; fprops.
      ue.
    rewrite drf => x xdf; bw.
    ex_middle nx; elim (h x); move: xdf; srw; ee.
  move: lt;rewrite -hf -hg eq; move => [_ ne]; elim ne; apply refl_equal.
have [l [lp [ln Vl]]]: exists l,
    (forall i : Set, l <c i -> i <c n -> V i f = V i g) &
      l <c n & (V l f) <> (V l g).
  set z:= Zo Bnat (fun l => l <c n & V l f <> V l g).
  have nz: nonempty z.
    move: pC => [l lp]; exists l; rewrite /z Z_rw; split=>//.
    move:lp => [[lp _] _]; Bnat_tac.
  have pa: (forall a, inc a z -> is_cardinal a).
    move=> a; rewrite /z Z_rw; move=> [aB _ ]; fprops.
  move: (wordering_cardinal_le_pr pa) => [sr wor].
  move: (worder_total wor) => tor.
  have zc: sub z (interval_co_0a n).
    move=> t; rewrite /z Z_rw; srw; intuition.
  have fsz: finite_set z.
    apply (sub_finite_set zc); red; rewrite cardinal_interval_co_0a1 //.
    by rewrite -inc_Bnat.
  have sw: sub z (substrate (graph_on cardinal_le z)) by rewrite sr; fprops.
  move: (finite_subset_torder_greatest tor fsz sw nz)=> [l []].
  move: wor => [or _].
  aw; move=> lz lp; exists l.
  move: (lz); rewrite /z Z_rw; move => [lB [ln Vl]]; ee.
  move=> i li lin.
  have iB: inc i Bnat by move: lin => [lin _ ]; Bnat_tac.
  case (equal_or_not (V i f) (V i g)) => // h.
  have iz: inc i z by rewrite /z Z_rw; ee.
  move: (lp _ iz); rewrite (iorder_gle _ iz lz) /gle graph_on_rw1.
  move=> [_ [ _ il]]; co_tac.
exists l; exists n; ee.
  have p1: is_cardinal (V l f).
    move: ef =>[bB [bK [b2 [fgf [df vf]]]]].
    have ldf: (inc l (domain f)) by rewrite df; srw; co_tac.
    by move: (vf _ ldf)=> [[res _] _ ].
  have p2: is_cardinal (V l g).
    move: eg =>[bB [bK [b2 [fgf [df vf]]]]].
    have ldf: (inc l (domain g)) by rewrite df; srw; co_tac.
    by move: (vf _ ldf)=> [[res _] _ ].
  case (cardinal_le_total_order2 p2 p1) => //.
  move=> h.
  have p3: (V l g) <c (V l f) by split; intuition.
  have lp' : forall i, cardinal_lt l i -> i <c n -> V i g = V i f.
    by move=> i i1 i2; symmetry;move: (lp _ i1 i2).
  move: (is_expansion_prop12 ng nf ln pB pA lp' eg ef p3) => [p4 _].
  co_tac.
Qed.

Lemma is_expansion_exists1: forall a b k,
  inc b Bnat -> 1c <c b -> inc k Bnat ->
  inc a Bnat -> a <c (b ^c k) ->
  exists f, (is_expansion f b k & expansion_value f b = a).
Proof.
move=> a b k bB b1 kB; move: k kB a.
apply :cardinal_c_induction.
  rewrite power_x_0 => c cB c1.
  exists (L emptyset (fun _ => card_zero)); split;ee;bw.
    red; ee.
      by bw; rewrite - emptyset_interval_00.
    bw; move=> i ie; elim (emptyset_pr ie).
  rewrite /expansion_value; bw; rewrite trivial_cardinal_sum; bw.
  move: c1; rewrite -succ_zero (lt_is_le_succ _ inc0_Bnat) => aux.
  by symmetry; apply zero_smallest2.
move=> n nB pn c cB cp.
set (b0:= b ^c n).
have BN: (inc b0 Bnat) by rewrite /b0; hprops.
have Bz: (b0 <> card_zero).
  have aux:(0c <c b) by move: zero_le_one => aux'; co_tac.
  have cn: (is_cardinal n) by fprops.
  have cb: is_cardinal b by fprops.
  have bnz: b<> card_zero by move: aux => []; intuition.
  apply (non_zero_apowb cb cn bnz).
move: (division_exists cB BN Bz) => [q [r [qB [rB [aux lt]]]]].
rewrite /b0 in lt.
move: (pn _ rB lt) => [f [ise ev]].
have p1: (b0 *c q) <c (b0 *c b).
  have p2: ((b0 *c q) <=c c).
    rewrite aux; apply sum_increasing3; fprops.
  have p3: (inc (b0 *c q) Bnat) by hprops.
  rewrite /b0 - pow_succ -/b0; fprops.
  co_tac.
move: (mult_lt_simplifiable BN qB bB Bz p1) => qb.
move: (is_expansion_prop8 ise qB qb).
set F:= (L _ _).
move=> [s1 s2]; exists F; ee.
rewrite s2 ev aux card_plusC /b0; apply refl_equal.
Qed.

Lemma is_expansion_exists: forall a b, inc a Bnat -> inc b Bnat ->
  1c <c b -> exists k, exists f,
    (is_expansion f b k & expansion_value f b = a).
Proof.
move=> a b aB bB b1; exists a.
apply is_expansion_exists1 => //.
apply lt_a_power_b_a => //.
Qed.

EIII-5-8 Combinatorial analysis


Theorem shepherd_principle: forall f c, is_function f ->
  (forall x, inc x (target f) -> cardinal (inv_image_by_fun f (singleton x))=c)
  -> cardinal (source f) = (cardinal (target f)) *c c.
Proof.
move=> f c ff cc.
set (pa := L (target f) (fun z=> (inv_image_by_fun f (singleton z)))).
have up: (unionb pa = (source f)).
  set_extens x ; rewrite /pa unionb_rw /inv_image_by_fun; bw.
    move=> [y [yt]];bw; aw.
    move=> [u]; aw; move=> [uy Jg]; rewrite uy in Jg; graph_tac.
  move => xsf; exists (W x f).
  move: (inc_W_target ff xsf)=> Wt.
  split =>//; bw; aw; exists (W x f); aw; ee; apply W_pr3 =>//.
have md: (mutually_disjoint pa).
  apply mutually_disjoint_prop.
  rewrite/pa /inv_image_by_fun; bw=> i j y it jt; bw.
  aw; move=> [u [ui Jg]] [v [vi Jg']]; awi ui ; awi vi.
  have fgf: fgraph (graph f) by fprops.
  rewrite -ui -vi; apply (fgraph_pr fgf Jg Jg').
have fgp: (fgraph pa) by rewrite /pa; gprops.
move: (cardinal_sum_pr fgp).
set f' := L _ _.
have ->: f' =cst_graph (domain pa) c.
  rewrite /f';apply L_exten1; rewrite /pa; bw; move => x xtf; bw; apply cc=>//.
rewrite sum_of_same1.
have ->:c *c (domain pa) = (cardinal (target f)) *c c.
  rewrite card_multC /pa; bw.
  apply card_mult_pr2; fprops; symmetry;apply cardinal_of_cardinal; fprops.
move => <-.
rewrite -up; aw; apply equipotent_disjoint_union =>//;
 last (by apply disjoint_union_disjoint);
   rewrite /disjoint_union_fam; bw; gprops.
move=> i id; bw; fprops.
Qed.

Definition eqmod a b B:= a %%c B = b %%c B.

Lemma card_rem_prop: forall a b B, inc a Bnat -> inc b Bnat -> inc B Bnat ->
  B <> 0c -> eqmod ((B *c a) +c b) b B.
Proof.
move=> a b B aB bB BB Bz;
rewrite /eqmod.
move: (Bnat_division bB BB Bz).
set q := b %/c B; set r := b %%c B.
move=> [r1B [q1B [aeq r1p]]].
rewrite aeq card_plusA - cardinal_distrib_prod_sum3.
set q2:= (a +c q).
have q2B: inc q2 Bnat by rewrite /q2; apply BS_plus.
set A:= ((B *c q2) +cr).
have dp: (division_prop A B q2 r) by split;[apply refl_equal | done].
move: (Bnat_quorem_pr (BS_plus (BS_mult BB q2B) r1B) BB q2B r1B dp).
by move=> [_ h]; symmetry.
Qed.

Lemma card_plus_permute24: forall s1 s2 r1 r2,
  ((s1 +c s2) +c r1) +c r2 = (s1 +c r1) +c (s2 +c r2).
Proof.
move=> s1 s2 r1 r2.
rewrite (card_plusA (card_plus s1 r1) s2 r2).
rewrite -(card_plusA s1 s2 r1) (card_plusC s2 r1).
rewrite (card_plusA s1 r1 s2); done.
Qed.

Lemma card_rem_sum: forall a b B, inc a Bnat -> inc b Bnat -> inc B Bnat ->
  B <> 0c -> eqmod (a +c b) ((a %%c B) +c (b %%c B)) B.
Proof.
move=> a b B aB bB BB Bz.
move: (Bnat_division aB BB Bz) (Bnat_division bB BB Bz).
rewrite /division_prop.
set q1:= a %/cB; set q2:= b %/c B; set r1:= a %%c B; set r2:= b %%c B.
move=> [r1B [q1B [aeq r1p]]][r2B [q2B [beq r2p]]].
rewrite aeq beq.
set t := card_plus _ _.
have ->: t= (B *c (q1 +c q2)) +c (r1 +c r2).
  rewrite cardinal_distrib_prod_sum3 card_plusA.
  set (s1:= B *c q1); set (s2:= B *c q2).
  by rewrite (card_plus_permute24 s1 s2 r1 r2).
apply: (card_rem_prop (BS_plus q1B q2B) (BS_plus r1B r2B) BB Bz).
Qed.

Lemma card_rem_mult: forall a b B, inc a Bnat -> inc b Bnat -> inc B Bnat ->
  B <> 0c -> eqmod (a *c b) ((a %%c B) *c (b %%c B)) B.
Proof.
move=> a b B aB bB BB Bz.
move: (Bnat_division aB BB Bz) (Bnat_division bB BB Bz).
set q1:= a %/cB; set q2:= b %/c B; set r1:= a %%c B; set r2:= b %%c B.
rewrite /division_prop.
move=> [r1B [q1B [aeq r1p]]][r2B [q2B [beq r2p]]].
rewrite aeq beq.
rewrite cardinal_distrib_prod_sum3(card_multC B q2) card_multA (card_multC _ B).
set v := ((B *c q1) +c r1) *c q2.
rewrite cardinal_distrib_prod_sum2 - card_multA card_plusA.
rewrite - cardinal_distrib_prod_sum3.
apply card_rem_prop;[ rewrite /v; hprops| hprops | exact | exact].
Qed.

Lemma eqmod_plus : forall a b a' b' B, inc a Bnat -> inc b Bnat ->
  inc a' Bnat -> inc b' Bnat -> inc B Bnat -> B <> 0c ->
  eqmod a a' B -> eqmod b b' B -> eqmod (a +c b) (a' +c b') B.
Proof.
move=> a b a' b' B aB bB a'B b'B BB Bnz e1 e2.
rewrite /eqmod in e1 e2 |- *.
by rewrite (card_rem_sum aB bB BB Bnz) (card_rem_sum a'B b'B BB Bnz) e1 e2.
Qed.

Lemma eqmod_mult : forall a b a' b' B, inc a Bnat -> inc b Bnat ->
  inc a' Bnat -> inc b' Bnat -> inc B Bnat -> B <> 0c ->
  eqmod a a' B -> eqmod b b' B -> eqmod (a *c b) (a' *c b') B.
Proof.
move=> a b a' b' B aB bB a'B b'B BB Bnz e1 e2.
rewrite /eqmod in e1 e2 |- *.
by rewrite (card_rem_mult aB bB BB Bnz) (card_rem_mult a'B b'B BB Bnz) e1 e2.
Qed.

Lemma eqmod_rem: forall a B, inc a Bnat -> inc B Bnat ->
  B <> 0c -> eqmod a (a %%c B) B.
Proof.
move=> a B aB BB Bz.
rewrite {1} (Bnat_div_pr aB BB); apply card_rem_prop; hprops.
Qed.

Lemma eqmod_succ : forall a a' B, inc a Bnat -> inc a' Bnat ->
  inc B Bnat -> B <> card_zero ->
  eqmod a a' B -> eqmod (succ a) (succ a') B.
Proof.
move=> a a' B aB a'B BB Bnz e1.
rewrite !cardinal_succ_pr4; fprops; apply eqmod_plus => //; fprops.
Qed.

Lemma eqmod_pow1 : forall a n B, inc a Bnat -> inc n Bnat ->
  inc B Bnat -> B <> 0c ->
  eqmod a 1c B -> eqmod (a ^c n) 1c B.
Proof.
move=> a n B aB nB BB Bnz h; move: n nB.
apply cardinal_c_induction.
  rewrite power_x_0 //.
move=> n nB h1.
rewrite (pow_succ _ nB).
move: (eqmod_mult (BS_pow aB nB) aB inc1_Bnat inc1_Bnat BB Bnz h1 h).
rewrite (one_unit_prodr); fprops.
Qed.

Lemma eqmod_pow2 : forall a b n B, inc a Bnat -> inc b Bnat -> inc n Bnat ->
  inc B Bnat -> B <> 0c ->
  eqmod a 1c B -> eqmod (b *c (a ^c n)) b B.
Proof.
move=> a b n B aB bB nB BB Bnz h.
move: (eqmod_pow1 aB nB BB Bnz h) => h2.
have aux: eqmod b b B by [].
move: (eqmod_mult bB (BS_pow aB nB) bB inc1_Bnat BB Bnz aux h2).
rewrite one_unit_prodr; fprops.
Qed.

Lemma eqmod_pow3 : forall f b k B, is_expansion f b k ->
  inc B Bnat -> B <> 0c -> eqmod b 1c B ->
  eqmod (expansion_value f b) (cardinal_sum f) B.
Proof.
move=> f b k B ep BB Bnz b1.
have kB: inc k Bnat by move: ep => [bB [kB _]].
move: k kB f ep.
set (p:= fun n => forall f, is_expansion f b n ->
  eqmod (expansion_value f b) (cardinal_sum f) B).
apply (cardinal_c_induction (r:=p)).
  move=> g [bB [_ [b2 [fgf [df _]]]]].
  have dge: domain g = emptyset by move: df; rewrite emptyset_interval_00.
  rewrite /expansion_value dge !trivial_cardinal_sum //; bw.
move=> n nB pn g eg.
have cn: (is_cardinal n) by fprops.
rewrite (is_expansion_prop7 eg cn).
move: (is_expansion_prop5 eg cn) => er.
move: (eg) => [bB[_[_ [fgg [dg _]]]]].
move: (BS_succ nB) => snB.
have si: sub (interval_co_0a n) (domain g).
  rewrite dg; apply (interval_co_0a_increasing nB).
have ->: (cardinal_sum g =
   (cardinal_sum (restr g (interval_co_0a n))) +c (V n g)).
  pose h w := 0c +c (V w g).
  have hp: forall x, inc x (domain g) -> h x = V x g.
    move=> x; rewrite dg (interval_co_0a_pr2 _ snB) => p1.
    move: (is_expansion_prop1 eg p1) => p2.
    rewrite /h; aw; fprops.
  move: (induction_on_sum h nB); simpl.
  have ->: h n = V n g by apply hp; rewrite dg; srw; fprops.
  have ->: (L (interval_co_0a (succ n)) h = g).
    apply fgraph_exten; gprops; first by symmetry;bw.
    bw; move=> x xd; bw; apply hp; ue.
  have ->: (L (interval_co_0a n) h = (restr g (interval_co_0a n))).
    apply fgraph_exten; [ gprops | fprops | | ].
      by bw; rewrite (restr_domain1 fgg si).
    bw; move=> x xb; bw; apply (hp _ (si _ xb)).
  by move => ->.
have p1: inc (V n g) Bnat by apply (is_expansion_prop1 eg (is_lt_succ nB)).
have p2: inc ((V n g) *c (b ^c n)) Bnat.
   apply (BS_mult p1 (BS_pow bB nB)).
have p3: inc (cardinal_sum (restr g (interval_co_0a n))) Bnat.
  apply finite_sum_finite_aux; last by exact.
  red; ee.
    move=> i; rewrite dg (interval_co_0a_pr2 _ snB).
    apply (is_expansion_prop1 eg).
  by red; rewrite dg (cardinal_interval_co_0a1 snB) - inc_Bnat.
apply (eqmod_plus (is_expansion_prop3 er) p2 p3 p1 BB Bnz (pn _ er)).
apply (eqmod_pow2 bB p1 nB BB Bnz b1).
Qed.

Definition card_five := succ card_four.
Definition card_ten := card_five +c card_five.
Notation "10 'c'" := card_ten.

Lemma inc5_Bnat : inc card_five Bnat.
Proof. apply (BS_succ inc4_Bnat). Qed.

Lemma inc10_Bnat : inc 10c Bnat.
Proof. apply (BS_plus inc5_Bnat inc5_Bnat). Qed.

Lemma card_plus_3_2: 3c +c 2c = card_five.
Proof.
rewrite /card_five (Bsucc_rw inc4_Bnat) /card_four (Bsucc_rw inc3_Bnat).
by rewrite - card_plusA card_two_pr.
Qed.

Lemma card_mult_3_3: 10c = succ (3c *c 3c).
Proof.
have aux: forall n, is_cardinal n -> (n +c n) +c n = n *c 3c.
  move=> n cn; rewrite /card_three (Bsucc_rw inc2_Bnat).
  rewrite cardinal_distrib_prod_sum3.
  by rewrite (one_unit_prodr cn) card_multC two_times_n.
rewrite /10c - card_plus_3_2 - card_plus_permute24.
set t:= (3c +c 3c).
rewrite - card_plusA -{2} succ_one.
rewrite (plus_via_succ _ inc1_Bnat) - (Bsucc_rw inc2_Bnat) -/card_three.
rewrite (plus_via_succ _ inc3_Bnat); apply f_equal.
rewrite /t; apply aux; apply (Bnat_is_cardinal inc3_Bnat).
Qed.

Lemma card_mult_10_3: eqmod 10c 1c 3c.
Proof.
rewrite card_mult_3_3 (Bsucc_rw (BS_mult inc3_Bnat inc3_Bnat)).
apply (card_rem_prop inc3_Bnat inc1_Bnat inc3_Bnat).
rewrite /card_three; apply succ_nonzero.
Qed.

Definition is_base_ten_expansion f k :=
  inc k Bnat & fgraph f & domain f = interval_co_0a k &
  forall i, inc i (domain f) -> (V i f) <c 10c.

Lemma divisibiliy_by_three : forall f k, is_base_ten_expansion f k ->
  let g:= (L (domain f) (fun i=> (V i f) *c (10c ^c i))) in
  eqmod (cardinal_sum g) (cardinal_sum f) 3c.
Proof.
move=> f k [p1 [p2 [p3 p4]]].
have C30: card_three <> card_zero by rewrite /card_three; apply succ_nonzero.
have ep: is_expansion f 10c k.
  red; ee; first by apply inc10_Bnat.
  move: (BS_mult inc3_Bnat inc3_Bnat) => i9B.
  rewrite card_mult_3_3;srw.
  apply one_small_cardinal; first by fprops.
  apply (zero_cardinal_product2 C30 C30).
exact (eqmod_pow3 ep inc3_Bnat C30 card_mult_10_3).
Qed.


Lemma integer_induction0: forall h a, exists_unique
   (fun f =>
   source f = Bnat & surjection f &
    W 0c f = a
    & forall n, inc n Bnat -> W (succ n) f = h n (W n f)).
Proof.
split.
set (M := fun u => cardinal (source u)).
set (p:= fun u => Yo(M u = 0c) a (h (predc (M u))(W (predc (M u)) u))).
move: Bnat_order_worder => wo.
move: (transfinite_definition p wo) => [[f df] _].
move: df; rewrite /transfinite_def Bnat_order_substrate
    /restriction_to_segment /p.
move=> [sf [srcf fp]].
have p1: forall a, inc a Bnat -> M (restriction1 f (segment Bnat_order a))= a.
  move=> b bB; rewrite /restriction1 /M; aw.
  rewrite segment_Bnat_order // cardinal_interval_co_0a1 //.
exists f; ee.
  rewrite (fp _ inc0_Bnat)(p1 _ inc0_Bnat) Y_if_rw //.
  move=> n nB.
  have sn: (inc (succ n) Bnat) by hprops.
  have sp: succ n <> 0c by apply succ_nonzero.
  rewrite (fp _ sn) (p1 _ sn) Y_if_not_rw // predc_pr1.
  rewrite restriction1_W //; first by fct_tac.
      rewrite srcf segment_Bnat_order //.
      apply sub_interval_co_0a_Bnat.
    rewrite segment_Bnat_order // interval_co_0a_pr3; fprops.
  fprops.
move=> x y [sx [sjx [x0 xs]]][sy [sjy [y0 ys]]].
apply function_exten4=>//; first by ue.
rewrite sx; apply cardinal_c_induction; first by ue.
by move=> n nB eq; rewrite (xs _ nB) (ys _ nB) eq.
Qed.

Lemma integer_induction: forall s a, exists_unique (fun f =>
  source f = Bnat & surjection f & W 0c f = a &
  forall n, inc n Bnat -> W (succ n) f = s (W n f)).
Proof.
move=> s a; set (h:= fun _ x:Set => s x).
move: (integer_induction0 h a)=> [[f fp] u].
split.
   exists f; auto.
move=> *; apply u=>//.
Qed.

Definition induction_defined0 h a := choose(fun f=>
  source f = Bnat & surjection f & W card_zero f = a &
  forall n, inc n Bnat -> W (succ n) f = h n (W n f)).

Lemma induction_defined_pr0: forall h a,
  let f := induction_defined0 h a in
     source f = Bnat & surjection f & W 0c f = a &
     forall n, inc n Bnat -> W (succ n) f = h n (W n f).
Proof.
move=> h a.
move:(integer_induction0 h a) => [e _].
rewrite /induction_defined0; apply choose_pr=> //.
Qed.

Definition induction_term s a := fun n => W n (induction_defined0 s a).

Lemma induction_term0: forall s a,
  induction_term s a 0c = a.
Proof.
move=> s a.
move: (induction_defined_pr0 s a)=> [sf [sjf [w0 ws]]].
rewrite /induction_term //.
Qed.

Lemma induction_terms: forall s a n,
  inc n Bnat ->
  induction_term s a (succ n) = s n (induction_term s a n).
Proof.
move=> s a.
move: (induction_defined_pr0 s a)=> [sf [sjf [w0 ws]]].
rewrite /induction_term //.
Qed.

Definition factorial n :=
  cardinal_prod (L (interval_co_0a n) succ).

Lemma factorial_succ: forall n, inc n Bnat ->
  factorial (succ n) = (factorial n) *c (succ n).
Proof.
by move=> b nB; rewrite (induction_on_prod succ nB).
Qed.

Lemma factorial0: factorial 0c = 1c.
Proof.
rewrite /factorial emptyset_interval_00 trivial_cardinal_prod =>//.
  gprops.
bw.
Qed.

Lemma factorial1: factorial 1c = 1c.
Proof.
move: inc0_Bnat => zb.
rewrite -succ_zero factorial_succ // factorial0; aw; apply succ_is_cardinal.
Qed.

Lemma factorial2: factorial 2c = 2c.
Proof.
rewrite -succ_one.
rewrite (factorial_succ inc1_Bnat) factorial1 succ_one;aw; fprops.
Qed.

Lemma factorial_nonzero: forall n, inc n Bnat -> factorial n <> 0c.
Proof.
apply cardinal_c_induction.
  rewrite factorial0.
  move: card_one_not_zero; intuition.
move=> m mB u; rewrite (factorial_succ mB).
apply zero_cardinal_product2 =>//.
apply succ_nonzero.
Qed.

Lemma BS_factorial: forall n, inc n Bnat -> inc (factorial n) Bnat.
Proof.
move=> n nB.
apply (cardinal_c_induction (r:=fun n=> inc (factorial n) Bnat)) =>//.
  rewrite factorial0; fprops.
move=> m mB u; rewrite (factorial_succ mB); hprops.
Qed.

Hint Resolve factorial_nonzero: fprops.

Lemma factorial_prop: forall f, f 0c = 1c ->
  (forall n, inc n Bnat -> f (succ n) = (f n) *c (succ n)) ->
  forall x, inc x Bnat -> f x = factorial x.
Proof.
move=> f fz fp; apply cardinal_c_induction.
  rewrite factorial0 //.
move=> m mB u; rewrite (factorial_succ mB) (fp _ mB) u // .
Qed.

Lemma factorial_induction: forall n, inc n Bnat ->
 factorial n = induction_term (fun a b => b *c(succ a)) 1c n.
Proof.
apply cardinal_c_induction.
  by rewrite factorial0 induction_term0.
by move=> m mB h; rewrite (factorial_succ mB) (induction_terms _ _ mB) h.
Qed.

Lemma quotient_of_factorials: forall a b,
  inc a Bnat -> inc b Bnat -> b <=c a ->
  (factorial b) %|c (factorial a).
Proof.
move=> a b aB bB ab; rewrite -(card_sub_rpr ab).
move: (BS_sub aB bB).
set (c := a -c b);rewrite card_plusC.
generalize c; apply cardinal_c_induction.
  rewrite (Bnat0_unit_sumr bB).
  apply BNdivides_itself.
   apply BS_factorial =>//.
   apply factorial_nonzero =>//.
move => n nB.
rewrite (plus_via_succ _ nB).
have p1: inc (b +c n) Bnat by hprops.
rewrite factorial_succ //;apply BNdivides_trans2; hprops.
Qed.

Lemma quotient_of_factorials1: forall a b,
  inc a Bnat -> inc b Bnat -> b <=c a ->
  (factorial (a -c b)) %|c (factorial a).
Proof.
move=> a b aB bB leab; apply (quotient_of_factorials aB).
  hprops.
apply sub_le_symmetry=>//.
Qed.

Definition number_of_injections b a :=
  (factorial a) %/c (factorial (a -c b)).

Lemma number_of_injections_pr: forall a b,
  inc a Bnat -> inc b Bnat -> b <=c a ->
  (number_of_injections b a) *c (factorial (a -c b)) = factorial a.
Proof.
rewrite/number_of_injections.
move=> a b aB bB leba; rewrite card_multC.
symmetry;apply BNdivides_pr3=>//.
apply quotient_of_factorials1 =>//.
Qed.

Lemma number_of_injections_int: forall a b,
  inc a Bnat -> inc b Bnat ->
  inc (number_of_injections b a) Bnat.
Proof.
rewrite/number_of_injections.
move=> a b aB bB.
apply (BS_quo (BS_factorial aB) (BS_factorial (BS_sub aB bB))).
Qed.

Lemma number_of_injections_base: forall a, inc a Bnat ->
  number_of_injections 0c a = 1c.
Proof.
rewrite /number_of_injections=> a aB.
rewrite minus_n_0C=> //.
apply (BNquo_itself (BS_factorial aB) (factorial_nonzero aB)).
Qed.

Lemma number_of_injections_rec: forall a b,
  inc a Bnat -> inc b Bnat -> b <c a ->
  (number_of_injections b a) *c (a -c b) =
  number_of_injections (succ b) a.
Proof.
move=> a b aB bB ltab.
have leab: cardinal_le b a by move: ltab => [le _].
move: (number_of_injections_pr aB bB leab).
set A:= number_of_injections b a.
have sB: inc (succ b) Bnat by hprops.
rewrite -lt_n_succ_le // in ltab.
rewrite - (number_of_injections_pr aB sB ltab).
set B:= number_of_injections (succ b) a.
have sb: inc (a -c (succ b)) Bnat by hprops.
rewrite (succ_sub aB bB ltab).
rewrite factorial_succ; last by fprops.
set w := (succ (card_sub a (succ b))).
have wB: inc w Bnat by rewrite /w; hprops.
set D:= factorial (a -c (succ b)).
rewrite (card_multC D _) card_multA.
apply mult_simplifiable_right.
      rewrite /D; apply BS_factorial =>//.
    apply BS_mult.
      rewrite /A; apply number_of_injections_int =>//.
    exact.
  rewrite /B; apply number_of_injections_int =>//.
apply factorial_nonzero =>//.
Qed.

Lemma cardinal_complement_image1: forall f, injection f ->
  (cardinal (complement (target f) (image_of_fun f)))
  +c (cardinal (source f))
  = cardinal (target f).
Proof.
move => f injf.
set A:= (image_of_fun f).
have ->: (cardinal (source f) = cardinal A).
  aw;exists (restriction_to_image f).
  split; first by apply (restriction_to_image_bijective injf).
  rewrite /restriction_to_image/restriction2; aw; ee.
symmetry; rewrite card_plusC.
set B:= complement (target f) A.
have p0: sub A (target f).
  move=> t; rewrite /A/image_of_fun; aw; move=> [x [xsf Jg]].
  have ff: is_function f by fct_tac.
  graph_tac.
have <-:union2 A B = target f.
  set_extens x; rewrite /B union2_rw; srw.
  case; [move=> xA; apply (p0 _ xA) | intuition].
  move=> xt; case (inc_or_not x A); intuition.
apply card_plus_pr1; fprops.
rewrite /B; apply disjoint_complement.
Qed.

Lemma cardinal_complement_image: forall f, injection f ->
  finite_set (target f) ->
  ((cardinal (source f)) <=c (cardinal (target f)) &
    cardinal (complement (target f) (image_of_fun f)) =
    (cardinal (target f)) -c (cardinal (source f))).
Proof.
move=> f injf.
move: (cardinal_complement_image1 injf).
rewrite /finite_set.
set (a:= cardinal (source f)).
set (b:= cardinal (target f)).
set (c:= cardinal (complement (target f) (image_of_fun f))).
rewrite -inc_Bnat.
move=> cab fb; move: (fb); rewrite - cab => fs.
have aB: inc a Bnat.
  have ca:is_cardinal a by rewrite /a; fprops.
  apply (Bnat_in_sum ca fs).
have cB: inc c Bnat.
  have ca:is_cardinal c by rewrite /c; fprops.
  apply (Bnat_in_sum2 ca fs).
split.
  rewrite card_plusC; apply Bsum_increasing3=> //.
symmetry; apply card_sub_pr1 =>//.
Qed.

Definition set_of_injections E F :=
  Zo (set_of_functions E F)(injection).

Lemma number_of_injections_prop: forall E F,
  finite_set F ->
  cardinal E <=c cardinal F ->
  cardinal (set_of_injections E F) =
    number_of_injections (cardinal E) (cardinal F).
Proof.
move=> E F fsF lce.
set (m:= cardinal F).
have fse: finite_set E by red; red in fsF; Bnat_tac.
move: lce.
have mB: inc m Bnat by rewrite /m inc_Bnat; red in fsF.
pose s E := (cardinal E) <=c m ->
   cardinal (set_of_injections E F) = number_of_injections (cardinal E) m.
apply (finite_set_induction0 (s:=s)) =>//.
  rewrite /s cardinal_emptyset number_of_injections_base //.
  move=> _.
  set f:= empty_function_tg F.
  have injf: injection f.
    apply bl_injective;[move=> t tv| move => t t' tv];elim (emptyset_pr tv).
  have ff: is_function f by fct_tac.
  have : (inc f (set_of_functions (source f) (target f))) by aw.
    rewrite {2 3} /f bl_source bl_target => aux.
  have ->: (set_of_injections emptyset F = singleton f).
   set_extens v; rewrite /set_of_injections Z_rw.
     move=> [vs _]; rewrite (small_set_of_functions_source vs aux); fprops.
   rewrite singleton_rw; move=> ->;ee.
  by rewrite cardinal_singleton.
clear fse E.
move=> E a sE naE; rewrite /s.
rewrite (cardinal_succ_pr naE).
set (n := cardinal E).
move=> lem.
have snB: inc (succ n) Bnat by Bnat_tac.
have nB: inc n Bnat by apply BS_nsucc; rewrite /n; fprops.
have ltnm: n <c m by rewrite -lt_n_succ_le.
rewrite -(number_of_injections_rec mB nB ltnm).
have lenm: n <=c m by move: ltnm => [h _].
move: (sE lenm); rewrite -/n; move=> <-.
clear sE.
set (G2:= set_of_injections (tack_on E a) F).
set (G1:= set_of_injections E F).
set rf:=BL (restriction ^~ E) G2 G1.
have <-: (source rf = G2) by rewrite /rf; aw.
have <-: (target rf = G1) by rewrite /rf; aw.
have ta: transf_axioms (restriction ^~ E) G2 G1.
  move=> t; rewrite /G1/G2 /set_of_injections Z_rw Z_rw.
  rewrite ! set_of_functions_rw; move=> [[ft [st tt]] it].
  have sat: sub E (source t) by rewrite st; fprops.
  have fr: is_function (restriction t E) by apply restriction_function.
  ee; rewrite /restriction; aw.
  split=>//; aw; move=> x y xE yE; rewrite restriction_W // restriction_W //.
  move: it=> [_ it]; apply it; apply sat=>//.
have fr: is_function rf by apply bl_function.
apply shepherd_principle =>//.
move=> x.
set(K:= inv_image_by_fun rf (singleton x)).
rewrite /rf /G1 /set_of_injections bl_target Z_rw set_of_functions_rw.
move=> [[fx [sxE txF]] injx].
have fst: finite_set (target x) by red;rewrite txF;fprops.
move:(cardinal_complement_image injx fst)=> [_].
set (C:= complement F (image_of_fun x)).
rewrite sxE txF -/n -/m -/C => <-.
set (val:= BL (W a) K C).
have tav:(transf_axioms (W a) K C).
  rewrite /C /K/inv_image_by_fun /rf /BL=> t; srw;aw;move=> [u]; aw.
  move=> [uz [z [zG2 h]]]; rewrite uz in h.
  move: zG2; rewrite /G2 /set_of_injections Z_rw; aw.
  move=> [[fz [sz tgz]] injz].
  have asz: inc a (source z) by rewrite sz; fprops.
  rewrite -(pr1_def h); split.
    rewrite -tgz;apply inc_W_target=>//.
  rewrite /image_of_fun; aw; move=> [y [ysx JG]].
  rewrite sxE in ysx.
  have ysz: inc y (source z) by rewrite sz; fprops.
  have: (W a z = W y x) by graph_tac.
  have Ez: sub E (source z) by rewrite sz; fprops.
  rewrite -(pr2_def h) restriction_W // =>sW.
  move: injz=> [ _ hi]; rewrite - (hi _ _ asz ysz sW) in ysx; contradiction.
have fv: is_function val by rewrite /val; apply bl_function.
have sv: (source val = K) by rewrite /val; aw.
have tv: (target val = C) by rewrite /val; aw.
aw; exists val; ee; apply bl_bijective =>//.
  move=> u v; rewrite /K/inv_image_by_fun;aw.
  move=> [u' [usx Jg1]][v' [vsx Jg2]].
  move: usx vsx; aw; move=> ux vx; rewrite ux in Jg1; rewrite vx in Jg2.
  have uG2:inc u G2 by move:(inc_pr1graph_source fr Jg1); rewrite /rf; aw.
  have vG2:inc v G2 by move:(inc_pr1graph_source fr Jg2); rewrite /rf; aw.
  move: (W_pr fr Jg1) (W_pr fr Jg2); rewrite /rf bl_W // bl_W//.
  move: uG2 vG2; rewrite /G2 /set_of_injections Z_rw Z_rw;aw.
  move=> [[fu [su tu]] inju][[fvv [svv tvv]] injv] ru rv Wa.
  apply function_exten=>//; try ue.
  move=> t; rewrite su; aw; case; last by move=> ->.
  move=> tE.
  transitivity (W t x).
    rewrite -ru restriction_W //; rewrite su; fprops.
  rewrite -rv restriction_W //; rewrite svv; fprops.
move=> y; rewrite /C/K; srw; move=> [yF nyi].
have nasc: ~ (inc a (source x)) by ue.
set (f:= tack_on_f x a y).
have ff: is_function f by rewrite /f; apply tack_on_function=>//.
have sf: source f = tack_on E a by rewrite /f /tack_on_f; aw; ue.
have tf: target f = F .
   rewrite /f /tack_on_f; aw; rewrite txF; apply tack_on_when_inc =>//.
have injf: injection f.
  split =>//; move=> u v; rewrite sf /f;aw; case => xsf; case =>ysf.
  - rewrite ! tack_on_W_in //; try ue; move: injx => [_ injx]; apply injx; ue.
  - rewrite -sxE in xsf.
    rewrite ysf tack_on_W_out // tack_on_W_in //.
    move=> eq; rewrite -eq in nyi; elim nyi.
    rewrite /image_of_fun; aw; ex_tac; apply W_pr3 =>//.
  - rewrite -sxE in ysf.
    rewrite xsf tack_on_W_out // tack_on_W_in //.
    move=> eq; rewrite eq in nyi; elim nyi.
    rewrite /image_of_fun; aw; ex_tac; apply W_pr3 =>//.
  - by rewrite xsf ysf.
have rFE:(restriction f E = x).
  apply function_exten=>//.
  - apply restriction_function =>//; ue.
  - by rewrite / restriction; aw.
  - rewrite /restriction; aw; ue.
  - rewrite /restriction; aw.
    have Es: sub E (source (tack_on_f x a y))
      by rewrite /tack_on_f sxE; aw; fprops.
    move=> u uE; rewrite restriction_W // /f tack_on_W_in//; ue.
have fG2: inc f G2.
  rewrite /G2/set_of_injections; apply Z_inc=>//; aw; ee.
have fi:inc f (inv_image_by_fun rf (singleton x)).
  rewrite /inv_image_by_fun; aw; exists x; ee.
  have <-: (W f rf = x) by rewrite /rf bl_W//.
  apply W_pr3=>//; rewrite /rf;aw.
ex_tac.
by rewrite /f tack_on_W_out.
Qed.

Lemma number_of_permutations: forall E, finite_set E ->
  cardinal (set_of_permutations E) = (factorial (cardinal E)).
Proof.
move=> E fsE.
have ->: (set_of_permutations E = set_of_injections E E).
  rewrite /set_of_permutations/set_of_injections.
  set_extens x; rewrite Z_rw Z_rw.
    rewrite /bijection; ee.
  move=> [xE ix]; ee.
  move: xE; aw; move=> [fx [sx tx]].
  apply bijective_if_same_finite_c_inj =>//; [ rewrite sx tx// | ue ].
have aux: (cardinal E) <=c (cardinal E) by fprops.
rewrite (number_of_injections_prop fsE aux).
have cE: inc (cardinal E) Bnat by rewrite inc_Bnat =>//.
rewrite /number_of_injections minus_n_nC // factorial0 BN_quo_one //.
by apply BS_factorial.
Qed.

Definition partition_with_pi_elements p E f :=
  domain f = domain p &
  (forall i, inc i (domain p) -> cardinal (V i f) = V i p) &
  partition_fam f E.

Definition set_of_partitions p E :=
  Zo(set_of_gfunctions (domain p) (powerset E))
    (partition_with_pi_elements p E).

Lemma set_of_partitions_rw: forall p E f,
  inc f (set_of_partitions p E) = partition_with_pi_elements p E f.
Proof.
move=> p E f; rewrite /set_of_partitions;apply iff_eq.
  rewrite Z_rw; intuition.
move=> h; apply Z_inc=>//.
move: h => [p2 [_ [p1 [_ p6]]]].
set (g := corresp (domain p) (powerset E) f).
have ->: (domain p = source g) by rewrite /g -p2; aw.
have ->: (powerset E = target g) by rewrite /g; aw.
have ->: (f = graph g) by rewrite /g; aw.
apply inc_set_of_gfunctions.
rewrite /g; apply is_function_pr =>//.
move=> t; rewrite (frange_inc_rw _ p1).
move=> [x [xdf tV]].
aw;move => y; rewrite -p6;srw; rewrite tV => yt.
ex_tac.
Qed.

Lemma fif_cardinal: forall i p,
  finite_int_fam p -> inc i (domain p) -> is_cardinal (V i p).
Proof. by move=> i p [fgp [alip fdp]]; fprops. Qed.

Lemma pip_prop0: forall p E f, partition_with_pi_elements p E f ->
  forall i, inc i (domain f) -> sub (V i f) E.
Proof.
move=> p E f [df [cV [fgf [duj ue]]]].
move=>i idp t tW; rewrite -ue; srw; ex_tac.
Qed.

Lemma number_of_partitions1: forall p E,
  finite_int_fam p -> cardinal_sum p = cardinal E ->
  nonempty (set_of_partitions p E).
Proof.
move=> p E fif.
rewrite /cardinal_sum cardinal_equipotent; move => [f [[bf sjf] [sf tg]]].
move: sf; rewrite /disjoint_union /disjoint_union_fam.
pose k i :=product (V i p) (singleton i).
move => sf.
have sx: forall i, inc i (domain p) ->sub (k i) (source f).
  move=> i idp t; rewrite sf unionb_rw; bw; exists i; ee; bw.
set g:= L (domain p) (fun i => image_by_fun f (k i)).
exists g; rewrite set_of_partitions_rw; red.
ee.
    rewrite /g; bw.
  move=> i idp; rewrite /g /k;bw.
  rewrite -{2} (cardinal_of_cardinal (fif_cardinal fif idp)).
  aw;eqsym;eqtrans (product (V i p) (singleton i)).
  apply equipotent_restriction1 =>//; apply sx =>//.
have ff: is_function f by fct_tac.
red; ee.
    rewrite /g;gprops.
  rewrite /g.
  apply mutually_disjoint_prop; bw.
  move=> i j y idp jdp; bw.
  move: (sx _ idp)(sx _ jdp)=> p1 p2.
  rewrite /k; aw;move =>[u [uki Wu]][v [vkj Wv]].
  have uv: u = v.
     move: bf => [_ injf].
     rewrite - Wv in Wu; apply injf =>//;[apply (p1 _ uki) | apply (p2 _ vkj)].
  move: uki vkj; rewrite uv /k.
  by aw; move=> [_[_ q1]] [_ [_ q2]]; rewrite -q1 -q2.
rewrite /g -tg.
set_extens x; srw; bw.
  move=> [y [yd]]; move: (sx _ yd) => aux; bw; aw ; move=> [u [uk]] <-.
  by apply inc_W_target =>//; apply aux.
move=> xtf; move: (surjective_pr2 sjf xtf)=> [y [ysf Wy]].
move: ysf; rewrite sf; srw; bw; move=> [z [zp]].
by bw; move=> ykz; exists z;ee; bw; aw; [ exists y; ee | apply sx].
Qed.

Definition set_of_partitions_aux f g:=
  L (domain f) (fun i => image_by_fun g (V i f)).

Lemma number_of_partitions3: forall p E f g,
  partition_with_pi_elements p E f -> inc g (set_of_permutations E) ->
  inc (set_of_partitions_aux f g) (set_of_partitions p E).
Proof.
move=> p E f g pip gp.
rewrite set_of_partitions_rw.
move: gp; rewrite /set_of_permutations Z_rw; aw;move => [[fg [sg tg]] [ig sjg]].
rewrite/ set_of_partitions_aux.
have Ha:forall u, inc u (domain f) -> sub (V u f) (source g).
  rewrite sg; apply (pip_prop0 pip).
move: pip => [df [cV [fgf [duj ue]]]].
red;bw; ee.
  rewrite df;move=> i idp; bw; rewrite - (cV _ idp) ; symmetry; aw.
  apply equipotent_restriction1 =>//; apply Ha =>//; ue.
red; ee.
  apply mutually_disjoint_prop=>//.
  bw; move=> i j y idp jdp.
  move: (Ha _ idp) (Ha _ jdp)=> s1 s2.
  bw; aw; move=> [u [ui Wu]][v [vi Wv]].
  have uv: u = v.
    move: ig => [_ ig]; rewrite -Wv in Wu;apply ig; fprops.
  red in duj; case (duj _ _ idp jdp)=> // ns.
  rewrite /W in ui vi; rewrite -uv in vi.
  red in ns; empty_tac1 u; apply intersection2_inc =>//.
rewrite -tg;set_extens x; srw;bw.
  move=> [y [ydp]]; move: (Ha _ ydp)=> s; bw; aw.
  by move=> [u [uW]] <-; apply inc_W_target => //; apply s.
move=> xtg; move: (surjective_pr2 sjg xtg)=> [y [ysf Wy]].
move: ysf; rewrite sg -ue; srw; move=> [t [tf etc]].
ex_tac; bw; aw; [ exists y; ee | by apply Ha].
Qed.

Lemma number_of_partitions4: forall p E f,
  finite_int_fam p -> cardinal_sum p = cardinal E ->
  partition_with_pi_elements p E f ->
  surjection (BL (set_of_partitions_aux f)
    (set_of_permutations E) (set_of_partitions p E)).
Proof.
move=> p E f fip spE [df [cV pfa]];move:(pfa)=> [fgf [duj ue]].
set phi:=BL _ _ _.
apply bl_surjective.
  move=> g; apply number_of_partitions3 =>//.
have Ha:finite_set E.
  by red;rewrite -spE -inc_Bnat; apply finite_sum_finite.
move=> y; rewrite set_of_partitions_rw; move=> [dy [cVy [fgy [dujy uey]]]].
pose ha i := equipotent_ex (V i f) (V i y).
have hap: forall i, inc i (domain p) ->
    (bijection (ha i) & source (ha i) = V i f & target (ha i) = V i y).
   move=> i idp; rewrite /ha; apply equipotent_ex_pr.
   by rewrite - cardinal_equipotent (cV _ idp) (cVy _ idp).
pose h i := BL (fun z=> W z (ha i)) (V i f) E.
have ta1:forall i, inc i (domain p) -> transf_axioms (fun z => W z (ha i))
    (V i f) E.
  move=> i idp z zW.
  move: (hap _ idp)=> [[[fh _] _] [sh th]].
  rewrite -sh in zW;move: (inc_W_target fh zW); rewrite th -uey.
  srw; move=> h1; exists i; rewrite dy; ee.
have fp1:(forall i, inc i (domain f) -> function_prop (h i) (V i f) E).
  rewrite df /h=> i idp; split.
  apply bl_function; apply ta1 =>//.
  rewrite bl_source bl_target; split => //.
rewrite -df in hap.
move: (extension_partition pfa fp1) => [[g [[fg [sg tg] extg]]] _].
have injg: injection g.
  split=> //; move=> u v.
  rewrite sg -ue; srw; move=> [u' [u'd Vu']][v' [v'd Vv']].
  have pu : W u g = W u (ha u').
    have ->: W u g = W u (h u').
      move: (extg _ u'd); rewrite /agrees_on; intuition.
    rewrite /h; aw; apply ta1; ue.
  have pv: W v g = W v (ha v').
    have ->: W v g = W v (h v').
      by move: (extg _ v'd); rewrite /agrees_on; intuition.
    rewrite /h; aw; apply ta1; ue.
  move=> eq; suff:u' = v'.
    move=> aux.
    move: pv; rewrite -aux -eq pu.
    move: (hap _ u'd) => [[[_ ih] _] [sa _]].
    rewrite sa in ih; apply ih =>//; ue.
  case (duj _ _ u'd v'd) =>// dj.
  have Wu: (inc (W u g) (V u' y)).
    rewrite pu; move: (hap _ u'd) => [[fh _] [sa ta]].
    rewrite -ta; apply inc_W_target;[ fct_tac | ue ].
  have Wv: (inc (W v g) (V v' y)).
    rewrite pv; move: (hap _ v'd) => [[fh _] [sa ta]].
    rewrite -ta; apply inc_W_target;[ fct_tac | ue ].
  red in dujy; rewrite dy -df in dujy.
  case (dujy _ _ u'd v'd) =>//.
  rewrite /disjoint; move=> h1; rewrite eq in Wu.
  empty_tac1 (W v g); apply intersection2_inc =>//.
have bx: (bijection g).
  apply bijective_if_same_finite_c_inj=>//; [ by rewrite sg tg | ue].
have pg: inc g (set_of_permutations E).
  rewrite /set_of_permutations; apply Z_inc=>//; aw;ee.
ex_tac.
symmetry.
rewrite /set_of_partitions_aux; apply fgraph_exten =>//; [ gprops | bw; ue|].
rewrite dy -df; move=> x xdf; bw.
move: (hap _ xdf) => [bha [sha tha]].
move: (fp1 _ xdf) (extg _ xdf) => [fh [sh th]] [p2a [_ p2c]].
have p3 : forall i, inc i (V x f) -> W i g = W i (ha x).
  move=> i ins; rewrite (p2c _ ins).
  by rewrite /h; aw; apply ta1 =>//; ue.
rewrite -tha.
symmetry; set_extens u; rewrite /image_by_fun; aw.
  move=> [i [ins Jg]].
  rewrite -(W_pr fg Jg) (p3 _ ins); apply inc_W_target =>//; [fct_tac| ue].
move: bha => [iha sjha] ut.
move: (surjective_pr2 sjha ut)=> [v [vs Wv]].
rewrite sha in vs; rewrite - Wv - (p3 _ vs).
by ex_tac; apply W_pr3 =>//; apply p2a.
Qed.

Lemma number_of_partitions5: forall p E f g h,
  finite_int_fam p -> cardinal_sum p = cardinal E ->
  partition_with_pi_elements p E f ->
  inc h (set_of_permutations E) -> inc g (set_of_permutations E) ->
  (set_of_partitions_aux f g = set_of_partitions_aux f h) =
  (forall i, inc i (domain p) ->
      image_by_fun ( (inverse_fun h) \co g) (V i f) = (V i f)).
Proof.
move=> p E f g h fip spE pip; move: (pip_prop0 pip)=> Ht.
move: pip => [df [cV [fgf [duj ue]]]].
rewrite /set_of_permutations Z_rw Z_rw; aw.
move=> [[fh [sh th]] bh] [[fg [sg tg]] bg].
rewrite /set_of_partitions_aux.
have fih: is_function (inverse_fun h) by apply bijective_inv_function.
have cih: (inverse_fun h) \coP g by red; ee; aw; ue.
have fcih : is_function ((inverse_fun h) \co g) by fct_tac.
rewrite -df; apply iff_eq.
  move=> eq i idp; move: (f_equal (V i) eq); bw.
  rewrite /image_by_fun/inverse_fun/compose; aw.
  move=> eql; set_extens x;rewrite image_by_graph_rw.
    move=> [y [yW]]; rewrite inc_compose_rw; move=> [_ [z []]].
    rewrite pr1_pair pr2_pair inverse_graph_pair => Jg Jg1.
    have : inc z (image_by_graph (graph g) (V i f)) by aw; ex_tac.
    rewrite eql; aw; move => [x' [x'W Jg2]].
    move: bh => [injh _]; rewrite (injective_pr3 injh Jg1 Jg2) //.
  move=> xW.
  have xsh: inc x (source h) by rewrite sh; apply (Ht _ idp).
  have : (inc (W x h) (image_by_graph (graph g) (V i f))).
    rewrite eql image_by_graph_rw; ex_tac; graph_tac.
  rewrite image_by_graph_rw;move=> [x' [x'W Jg]]; ex_tac.
  rewrite inc_compose_rw; split; first by fprops.
  exists (W x h); rewrite pr1_pair pr2_pair inverse_graph_pair.
  by split; last by graph_tac.
move=> aux; apply L_exten1=>//.
move=> i xdp; rewrite /image_by_fun.
move: (aux _ xdp) => aux'.
set_extens x; aw; move=> [y [yW Jg]].
  have xt:inc x (target h)
     by rewrite th -tg ;apply (inc_pr2graph_target fg Jg).
  move: (inc_pr1graph_source fg Jg) =>ys.
  move: bh=> [_ sjh]; move: (surjective_pr2 sjh xt)=> [z [zs Wz]].
  have p2: inc (J z x) (graph h) by rewrite -Wz;apply W_pr3 =>//.
  ex_tac; rewrite -(aux _ xdp) /image_by_fun/compose /inverse_fun.
  rewrite corresp_graph image_by_graph_rw; exists y; split; [exact |].
  rewrite corresp_graph inc_compose_rw; split; first by fprops.
  by exists x; rewrite pr1_pair pr2_pair inverse_graph_pair; split.
move: yW; rewrite -{1} aux'/image_by_fun/compose /inverse_fun.
rewrite corresp_graph image_by_graph_rw;move=> [u [uW]].
rewrite corresp_graph inc_compose_rw;move => [_[v]].
rewrite pr1_pair pr2_pair inverse_graph_pair.
move=> [J1g J2g]; ex_tac.
have fgh: fgraph (graph h) by fprops.
by rewrite (fgraph_pr fgh Jg J2g).
Qed.

Lemma number_of_partitions6: forall p E f h,
  finite_int_fam p -> cardinal_sum p = cardinal E ->
  partition_with_pi_elements p E f ->
  inc h (set_of_permutations E) ->
  transf_axioms (fun g=> L (domain p)(fun i=> (restriction2
    ((inverse_fun h) \co g)
    (V i f) (V i f))))
  (Zo (set_of_permutations E)
    (fun g => (set_of_partitions_aux f g = set_of_partitions_aux f h)))
  (productb (L (domain p)(fun i=> (set_of_permutations (V i f))))).
Proof.
move=> p E f h fip spE pip hp g; move: (pip_prop0 pip)=> Ht.
move: (pip) => [df [cV [fgf [duj ue]]]].
rewrite Z_rw; move=> [gp eql].
aw;last by gprops.
bw; ee; move=> i idp; bw.
rewrite (number_of_partitions5 fip spE pip hp gp) in eql.
move: (eql _ idp) => ic; clear eql.
move: hp gp; rewrite /set_of_permutations Z_rw Z_rw; aw.
move=> [[fh [sh th]] bh] [[fg [sg tg]] bg].
have fih: is_function (inverse_fun h) by apply bijective_inv_function.
have cih: (inverse_fun h) \coP g by red; ee; aw; ue.
have fcih : is_function ((inverse_fun h) \co g) by fct_tac.
set k:= restriction2 _ _ _.
have sk: source k = V i f by rewrite /k/restriction2; aw.
have tk: target k = V i f by rewrite /k/restriction2; aw.
rewrite df in Ht; move: (Ht _ idp) => Ht1.
have ra: (restriction2_axioms (compose (inverse_fun h) g) (V i f) (V i f)).
   red; aw; ee; ue.
have fk: (is_function k) by rewrite /k; apply restriction2_function.
apply Z_inc; first by aw;ee.
apply bijective_if_same_finite_c_inj.
    by rewrite sk tk.
  red;rewrite sk; move: fip=> [_[hp _ ]]; move: (hp _ idp).
  move: pip => [_ [pip _]];rewrite (pip _ idp); bw.
split=>// x y; rewrite sk => xsk ysk.
rewrite /k restriction2_W //restriction2_W //.
move: (Ht1 _ xsk) (Ht1 _ ysk); rewrite -sg => xsg ysg.
aw;move => eq.
move: bg=> [[_ ig] _ ]; apply ig =>//.
move: (inverse_bij_is_bij bh) => [[_ iih] _ ].
awi iih; apply iih =>//; try (rewrite th -tg; apply inc_W_target =>//).
Qed.

Lemma number_of_partitions7: forall p E f h,
  finite_int_fam p -> cardinal_sum p = cardinal E ->
  partition_with_pi_elements p E f ->
  inc h (set_of_permutations E) ->
  bijection (BL (fun g=> L (domain p)(fun i=> (restriction2
    ((inverse_fun h) \co g)
    (V i f) (V i f))))
  (Zo (set_of_permutations E)
    (fun g => (set_of_partitions_aux f g = set_of_partitions_aux f h)))
  (productb (L (domain p)(fun i=> (set_of_permutations (V i f)))))).
Proof.
move=> p E f h fip spE pip hp; move: (pip_prop0 pip)=> sWE.
move: (pip) => [df [cV [fgf [duj ue]]]].
rewrite df in sWE.
set ww:=BL _ _ _.
move: (number_of_partitions6 fip spE pip hp) => ta.
move:(hp); rewrite /set_of_permutations Z_rw; aw; move=> [[fh [sh th]] bh].
have bih: (bijection (inverse_fun h)) by apply inverse_bij_is_bij.
have fw: is_function ww by rewrite /ww; apply bl_function.
move: (pip) => [_ [_ pfa]].
have sw: surjection ww.
  apply surjective_pr6 =>//.
  rewrite /ww bl_source bl_target => y.
  rewrite productb_rw ? L_domain; last (by gprops) ; move=> [fgy [dy yp]].
  set (ha := fun i=> V i y).
  have yp': (forall i, inc i (domain p) ->
    (bijection (ha i) & source (ha i) = V i f & target (ha i) = V i f)).
    move=> i idp; rewrite dy in yp; move: (yp _ idp); bw.
    rewrite /set_of_permutations Z_rw /ha set_of_functions_rw; intuition.
  pose hb i := BL (fun z=> W z (ha i)) (V i f) E.
  have ta1: (forall i, inc i (domain p) -> transf_axioms (fun z => W z (ha i))
    (V i f) E).
    move=> i idp z zW; move: (yp' _ idp)=> [bha [sha tha]].
    apply (sWE _ idp); rewrite -tha;apply inc_W_target; [fct_tac | ue].
  have Hw:forall i, inc i (domain f) -> function_prop (hb i) (V i f) E.
    rewrite df; move => i ip; rewrite /hb /function_prop bl_source bl_target.
    split; [ apply bl_function;apply ta1 =>// | ee].
  move: (extension_partition pfa Hw) => [[g [[fg [sg tg]] alag]] _].
  have ijg: injection g.
    split=>// u v; rewrite sg => uE vE.
    move: uE vE;rewrite -ue; srw; move=> [a [adf ua]] [b [bdf vb]].
    move: (alag _ adf)(alag _ bdf) => [p1 [p2 p3]] [q1 [q2 q3]].
    have ap: inc a (domain p) by ue.
    have bp: inc b (domain p) by ue.
    rewrite (p3 _ ua)(q3 _ vb).
    move : (ta1 _ ap) (ta1 _ bp) => ta2 ta3.
    rewrite /hb (bl_W ta2 ua) (bl_W ta3 vb).
    move: (yp' _ ap) (yp' _ bp)=> [[[_ iha] xx] [sha tha]] [yy [shb thb]].
    case (duj _ _ adf bdf).
      move=> ab; rewrite -ab in vb; rewrite -ab.
      apply iha; rewrite sha //.
    move=> h1 h2; rewrite /disjoint in h1; empty_tac1 (W u (ha a)).
    apply intersection2_inc.
     rewrite /W in tha;rewrite -tha; apply inc_W_target ; [ fct_tac | ue].
     rewrite /W in thb;rewrite h2 -thb; apply inc_W_target ; [ fct_tac | ue].
  have bg: (bijection g).
    apply bijective_if_same_finite_c_inj; last by exact.
       rewrite sg tg //.
    rewrite sg /finite_set - inc_Bnat -spE; apply finite_sum_finite =>//.
  have pa: forall i v, inc i(domain p) -> inc v (V i f) -> W v g = W v (ha i).
    rewrite -df;move=> i v idp; move: (alag _ idp).
    rewrite /agrees_on; move => [s1 [s2 aux]].
    rewrite {1} /W; move=> vW; move: (aux _ vW); rewrite /hb bl_W;
       [done | apply ta1;ue| done ].
  have pb:forall i v, inc i (domain p) -> inc v (V i f) -> inc (W v g)(V i f).
    move=> i v ip iv; rewrite (pa _ _ ip iv).
    move: (yp' _ ip) => [bha [sha tha]].
   rewrite -tha; apply inc_W_target; [fct_tac| ue].
  have chg: (h \coP g) by red; ee;ue.
  have fhg: is_function (h \co g) by fct_tac.
  have bhg: (bijection (h \co g)) by apply compose_bijective.
  set (t:= h \co g).
  have pt: inc t (set_of_permutations E).
    rewrite /set_of_permutations; apply Z_inc; aw;rewrite /t; aw; ee.
  set z:=Zo _ _.
  have xz:inc t z.
    rewrite /z; apply Z_inc =>//.
    rewrite /set_of_partitions_aux df; apply L_exten1=> //.
    move=> i idp; rewrite /image_by_fun /t/compose; set_extens u; aw.
      move=> [a [ai]]; aw; move=> [_ [b [Jg1 Jg2]]].
       by exists b; ee; move: (W_pr fg Jg1) => <-; apply pb.
    move=> [a [asW Jg]].
    have atg: inc a (target g) by rewrite tg; apply (sWE _ idp).
    move: (bg) => [_ sjg]; move: (surjective_pr2 sjg atg) => [b [bsg Wb]].
    move: (bsg); rewrite sg; rewrite -ue; srw.
    move => [j [jd bW]].
    rewrite -df in idp; case (duj _ _ idp jd).
      move=> eq; rewrite -eq in bW; exists b; split; first by exact.
      rewrite inc_compose_rw; split; first by fprops.
      rewrite pr1_pair pr2_pair; exists a; split; last by exact.
      rewrite -Wb; apply W_pr3 => //; rewrite sg; apply (sWE _ idp).
    move=> di; red in di; empty_tac1 a; apply intersection2_inc =>//.
    rewrite -df in pb; rewrite -Wb;apply (pb _ _ jd bW).
  ex_tac; rewrite (bl_W ta xz).
  have ->: ( (inverse_fun h) \co t = g).
   rewrite /t -(compose_assoc (composable_inv_f bh) chg).
   rewrite bij_left_inverse //.
     rewrite sh -tg compose_id_left //.
  apply fgraph_exten;bw; gprops.
  move=> x xdp; bw.
  move: (sWE _ xdp)=> Wf.
  have ra: (restriction2_axioms g (V x f) (V x f)).
   have ssg: sub (V x f) (source g) by rewrite sg.
   have stg: sub (V x f) (target g) by rewrite tg.
   red; ee; move=> u; aw;move=> [v [vW]] <-; apply (pb _ _ xdp vW).
  move : (yp' _ xdp); rewrite /ha; move=> [bha [sha tha]].
  apply function_exten.
          by apply restriction2_function.
        fct_tac.
      by rewrite /restriction2 sha corresp_source.
    by rewrite /restriction2 tha corresp_target.
  rewrite /restriction2;aw; move=> u uW; rewrite restriction2_W //.
  by rewrite (pa _ _ xdp uW).
split=>//; split =>//.
rewrite /ww bl_source=> x y xS yS; rewrite (bl_W ta xS) (bl_W ta yS).
move: xS yS; rewrite Z_rw Z_rw; move=> [xs eq1][ys eq2].
rewrite (number_of_partitions5 fip spE pip hp xs) in eq1.
rewrite (number_of_partitions5 fip spE pip hp ys) in eq2.
move: eq1 eq2.
set xi:= _ \co x; set yi:= _ \co y => eq1 eq2.
have shi: (source (inverse_fun h)= E) by aw.
have thi: (target (inverse_fun h)= E) by aw.
move: xs ys; rewrite /set_of_permutations Z_rw Z_rw ! set_of_functions_rw.
move=> [[fx [sx tx]] bx][[fy [sy ty]] bjy] eq.
have Hx: (inverse_fun h) \coP x by red; ee;[ fct_tac | ue].
have Hy: (inverse_fun h) \coP y by red; ee;[ fct_tac | ue].
suff: xi = yi.
  rewrite /xi /yi;move => eqi; move: (f_equal (compose h) eqi).
  move: (composable_f_inv bh) => co.
  rewrite -(compose_assoc co Hx) -(compose_assoc co Hy) (bij_right_inverse bh).
  by rewrite th -{1} tx -ty (compose_id_left fx) (compose_id_left fy).
have fxi: is_function xi by rewrite/xi; apply compose_function=>//.
have fyi: is_function yi by rewrite/xi; apply compose_function=>//.
have sxi: source xi = E by rewrite /xi compose_source.
have syi: source yi = E by rewrite /yi compose_source.
have txi: target xi = E by rewrite /xi compose_target.
have tyi: target yi = E by rewrite /yi compose_target.
apply (function_exten fxi fyi); [ by rewrite syi | by rewrite tyi| ].
rewrite sxi -ue;move=> u; srw; rewrite df; move=> [i [ip iW]].
move: (f_equal (V i) eq); bw => eq3.
have ra:restriction2_axioms xi (V i f) (V i f).
  red;ee; last (by rewrite eq1; fprops);rewrite ? sxi ? txi; apply (sWE _ ip).
have rb:restriction2_axioms yi (V i f) (V i f).
  red;ee; last (by rewrite eq2; fprops);rewrite ? syi ? tyi; apply (sWE _ ip).
move: (f_equal (W u) eq3); rewrite restriction2_W // restriction2_W //.
Qed.

Theorem number_of_partitions: forall p E,
  finite_int_fam p -> cardinal_sum p = cardinal E ->
  let num:= factorial (cardinal E) in
    let den := cardinal_prod (L (domain p) (fun z => factorial (V z p)))
      in (
        num = cardinal (set_of_partitions p E) *c den &
        inc num Bnat & inc den Bnat & den <> card_zero &
        finite_set (set_of_partitions p E)).
Proof.
move => p E fip spe num den.
move: (number_of_partitions1 fip spe)=> [h hp].
have Hb: finite_set E.
  by red; rewrite -spe -inc_Bnat; apply finite_sum_finite=> //.
have numB: inc num Bnat by rewrite /num; apply BS_factorial; aw; bw.
have ndenB: inc den Bnat.
  move: fip=> [p1 [p2 p3]].
  rewrite /den;apply finite_product_finite =>//.
  red; bw;ee; move=> i idp; bw; rewrite -inc_Bnat;apply BS_factorial.
  apply p2=>//.
have nd:den <> 0c.
  rewrite /den -zero_cardinal_product; bw; gprops.
  move: fip=> [p1 [p2 p3]].
  move=> i idp; bw; apply factorial_nonzero; apply p2=>//.
rewrite /finite_set.
set aux:= cardinal _.
suff: num = aux *c den.
  move => eql; ee.
  rewrite eql in numB.
  have ca: is_cardinal aux by rewrite /aux; fprops.
  rewrite -inc_Bnat; apply (Bnat_in_product ca nd numB).
move: hp; rewrite /set_of_partitions Z_rw; move => [hfpp pip].
move:(number_of_partitions4 fip spe pip).
set phi:= BL _ _ _; move=> sjphi.
have caux:(forall w, bijection w -> cardinal (source w) = cardinal (target w)).
  by move=> w bw; aw; exists w; ee.
suff phip: forall x, inc x (target phi) ->
    cardinal (inv_image_by_fun phi (singleton x)) = den.
  move: sjphi=> [fphi _].
  move:(shepherd_principle fphi phip).
  rewrite /phi bl_source bl_target number_of_permutations //.
move=> x xt.
move:(surjective_pr2 sjphi xt) => [y [ys yw]].
move: ys; rewrite {1}/ phi; aw => yp.
move: (number_of_partitions7 fip spe pip yp) => baux.
move: (caux _ baux); rewrite bl_source bl_target; clear baux.
set d:= cardinal (productb _).
have ->: d = den.
  rewrite/d/den/cardinal_prod; aw; apply equipotent_productb; gprops; bw.
  move=> i idp; bw.
  move: pip fip=>[p1 [p2 [p3 _]]][q1 [q2 q3]].
  rewrite - cardinal_equipotent number_of_permutations.
    rewrite (p2 _ idp).
    set (t := (factorial (V i p))).
    have ct: is_cardinal t by move: (BS_factorial (q2 _ idp)); fprops.
    by symmetry; apply cardinal_of_cardinal.
  red; rewrite p2 //; rewrite -inc_Bnat; apply q2 =>//.
move => <-. set z := Zo _ _.
have Ha:transf_axioms (fun g =>
     (set_of_partitions_aux h g))
    (set_of_permutations E) (set_of_partitions p E).
  move => g gp; apply number_of_partitions3 =>//.
apply f_equal.
rewrite /inv_image_by_fun /z.
have fphi: is_function phi by fct_tac.
set_extens t ; rewrite Z_rw; aw.
  move=>[u]; aw; move=> [ux Jg]; rewrite ux in Jg.
  move:(inc_pr1graph_source fphi Jg); rewrite /phi; aw => ts; ee.
  by move: (W_pr fphi Jg) yw; rewrite /phi; aw; move =>-> ->.
move=> [tp hh].
have ->: (x = W t phi) by rewrite -yw {1 2}/phi; aw.
ex_tac; graph_tac; rewrite /phi; aw.
Qed.

Theorem number_of_partitions_bis: forall p E,
  finite_int_fam p -> cardinal_sum p = cardinal E ->
  cardinal (set_of_partitions p E) =
  (factorial (cardinal E)) %/c
           (cardinal_prod (L (domain p) (fun z => factorial (V z p)))).
Proof.
move=> p E h1 h2.
move: (number_of_partitions h1 h2); simpl.
set (x1:= cardinal (set_of_partitions p E)).
set x2:= cardinal_prod _.
move=> [p1 [p2 [p3 [p4]]]].
rewrite /finite_set -inc_Bnat => p5.
rewrite p1 card_multC; symmetry; apply BNdivides_pr4 =>//.
Qed.

Lemma number_of_partitions_p2:
 forall E m p, inc m Bnat -> inc p Bnat -> cardinal E = (m +c p) ->
 let num := factorial (m +c p) in
 let den := (factorial m) *c (factorial p) in
 let x := cardinal (set_of_partitions (variantLc m p) E) in
   inc x Bnat & num = x *c den &
      inc num Bnat & inc den Bnat & den <> 0c.
Proof.
move=> E m p mB pB cE num den.
set P:= variantLc m p.
move=> x.
have h1: finite_int_fam P.
  rewrite /P;red; ee.
  by bw; move=> i it; try_lvariant it; rewrite - inc_Bnat.
  bw;rewrite -two_points_pr2; apply doubleton_finite.
have p2: cardinal_sum P = cardinal E.
  symmetry;rewrite cE /card_plus/cardinal_sum; fprops.
move: (number_of_partitions h1 p2); simpl.
set y:= cardinal_prod _.
have -> : den = y.
  rewrite /den/card_mult/y/P; bw.
  apply f_equal.
  apply fgraph_exten; fprops; gprops; bw.
  move=> xx xtp; try_lvariant xtp; fprops.
rewrite cE -/num /finite_set -/x.
move=> [r1 [r2 [r3 [r4 r5]]]]; split; [bw | split; [exact | ]].
split; [exact | split; [exact | exact]].
Qed.

Binomial coefficient

Definition binom n m :=
  V m (induction_term
      (fun _ T: Set => L Bnat (fun z => variant 0c 1c
          (V z T +c V (predc z) T) z))
       (L Bnat (variant 0c 1c 0c))
       n).

Lemma binom00: binom 0c 0c = 1c.
Proof.
by rewrite /binom induction_term0 /variant; bw; [ Ytac0 | fprops ].
Qed.

Lemma binom0Sm: forall m, inc m Bnat ->
  binom 0c (succ m) = 0c.
Proof.
move=> m mBnat.
rewrite /binom /variant induction_term0; bw; last by hprops.
have snz: succ m <> 0c by apply succ_nonzero.
by Ytac0.
Qed.

Lemma binomSn0: forall n, inc n Bnat-> binom (succ n) 0c = 1c.
Proof.
move=> n nB.
rewrite /binom /variant induction_terms //; bw; [by Ytac0 | hprops].
Qed.

Lemma binomSnSm: forall n m, inc n Bnat-> inc m Bnat ->
  binom (succ n) (succ m) = (binom n (succ m)) +c (binom n m).
Proof.
move=> n m nB mB.
rewrite /binom /variant induction_terms //.
set T := induction_term _ _.
have snz: succ m <> 0c by apply succ_nonzero.
bw; [Ytac0; rewrite predc_pr1 //; fprops | hprops ].
Qed.

Lemma BS_binom: forall n m, inc n Bnat -> inc m Bnat ->
  inc (binom n m) Bnat.
Proof.
move=> n m nB; move: n nB m.
apply: cardinal_c_induction.
  apply cardinal_c_induction.
      rewrite binom00=>//; fprops.
    move => n nB _; rewrite binom0Sm=>//; fprops.
move => n nB hrec; apply cardinal_c_induction.
  rewrite binomSn0=>//; fprops.
move => p pB bB.
rewrite binomSnSm=>//; hprops.
Qed.

Lemma binom_alt_pr: forall n m, inc n Bnat -> inc m Bnat ->
  (binom n m) *c (factorial m) *c (factorial (n -c m)) =
  Yo (m <=c n) (factorial n) 0c.
Proof.
pose Bi n m := (binom n m) *c
     (factorial m *c (factorial (card_sub n m))).
have ba4: forall n m, inc n Bnat -> inc m Bnat -> Bi (succ n) (succ m)=
    ((Bi n (succ m)) *c (Yo (n <=c m) 1c (n -c m)))
   +c (Bi n m *c succ m).
  move=> n m nB mB.
  set (x:= Bi n m); set (y:= Bi n (succ m)).
  rewrite /Bi binomSnSm //.
  rewrite card_sub_pr6 =>//.
  set aux := (factorial (succ m)) *c (factorial (n -c m)).
  rewrite cardinal_distrib_prod_sum2.
  set A:= (binom n (succ m)) *c aux.
  set (AA:= y *c (Yo (n <=c m) 1c (n -c m))).
  suff: A = AA.
    move=> ->; rewrite /x /Bi - card_multA; congr (_ +c (_ *c _)).
    rewrite (card_multC (factorial m) _) - card_multA -(factorial_succ mB).
    rewrite card_multC//.
  rewrite /A /AA /y /Bi; rewrite /aux - ! card_multA; congr (_ *c (_ *c _)).
  have nsmB: inc (n -c (succ m)) Bnat by hprops.
  case (p_or_not_p (n <=c m)) => le.
    have cf: is_cardinal (factorial (n -c (succ m))).
      apply Bnat_is_cardinal;apply (BS_factorial nsmB).
    Ytac0; rewrite (one_unit_prodr cf).
    have nel: ~ (succ m) <=c n.
      move: (is_lt_succ mB) => h1 el.
      have l3: (n <c succ m) by co_tac.
      co_tac.
    rewrite (card_sub_wrong nB (BS_succ mB) nel).
    case (equal_or_not n m) => eq.
     rewrite eq minus_n_nC =>//.
    rewrite card_sub_wrong; fprops; move=> el; elim eq; co_tac.
  Ytac0.
  have sm: (succ m) <=c n.
    case (cardinal_le_total_order2 (Bnat_is_cardinal nB)(Bnat_is_cardinal mB)).
      move => nle; contradiction.
    by rewrite (lt_n_succ_le mB nB).
  move: (succ_sub nB mB sm) => s.
  rewrite {1} s factorial_succ -? s; [ trivial | fprops].
move=> n m nB mB.
rewrite - card_multA; change (Bi n m = Yo (m <=c n) (factorial n) 0c).
move: n nB m mB; apply: cardinal_c_induction.
  move=> m mB.
  move: (BS_factorial mB) => fi.
  have fc: is_cardinal (factorial m) by fprops.
  have cs: (0c -c m) = 0c.
    case (equal_or_not m 0c) => h.
      rewrite h minus_n_nC; fprops.
    rewrite card_sub_wrong; fprops.
    by dneg h'; apply zero_smallest2.
  rewrite /Bi cs factorial0 (one_unit_prodr fc) {fi fc cs}.
  move: m mB; apply cardinal_c_induction.
    rewrite binom00 factorial0; aw; fprops.
    rewrite Y_if_rw; fprops.
  move=> p pB aux.
  have ns: ~ cardinal_le (succ p) 0c.
    rewrite lt_n_succ_le0 //; fprops.
    move=> h';apply (zero_smallest1 h').
  by Ytac0;rewrite binom0Sm //card_multC zero_prod_absorbing.
move=> N NB rN; apply cardinal_c_induction.
  move: (BS_succ NB) => sB.
  rewrite /Bi (binomSn0 NB) factorial0.
  rewrite (minus_n_0C sB).
  rewrite (one_unit_prodl (Bnat_is_cardinal (BS_factorial sB))).
  rewrite (one_unit_prodl (Bnat_is_cardinal (BS_factorial sB))).
  rewrite Y_if_rw//; fprops.
move => p pB bB.
have spN: inc (succ p) Bnat by hprops.
rewrite (ba4 _ _ NB pB).
rewrite (rN _ pB)(rN _ spN) factorial_succ //.
have Nc: is_cardinal N by fprops.
have psc: is_cardinal (succ p) by fprops.
have cp: is_cardinal p by fprops.
move: (is_lt_succ pB) => psp.
rewrite - (lt_n_succ_le1 cp NB).
case (cardinal_le_total_order2 psc Nc) => c1.
  have [pN npN]: cardinal_lt p N by co_tac.
  have yf:~ (cardinal_le N p) by dneg eql; co_tac.
  do 4 Ytac0; rewrite - cardinal_distrib_prod_sum3.
  by apply f_equal; rewrite (Bsucc_rw pB) (Bsucc_rw NB) card_plusA card_sub_rpr.
have r1: ~ ((succ p) <=c N) by move=> h; co_tac.
Ytac0; rewrite card_multC zero_prod_absorbing.
rewrite lt_is_le_succ in c1 =>//.
case (equal_or_not p N) => epN.
  have c2: cardinal_le p N by rewrite epN; fprops.
  Ytac0; Ytac0;rewrite epN.
  rewrite Bnat0_unit_suml =>//.
    apply (BS_mult (BS_factorial NB) (BS_succ NB)).
have c2: ~(p <=c N) by dneg c3; co_tac.
Ytac0; Ytac0; rewrite card_multC zero_prod_absorbing //.
by rewrite (Bnat0_unit_suml inc0_Bnat).
Qed.

Lemma binom_bad: forall n m, inc n Bnat -> inc m Bnat ->
  n <c m -> binom n m = 0c.
Proof.
move=> n m nB mB h; move: (binom_alt_pr nB mB).
set c := card_mult _ _.
have nz: ~ (m <=c n) by move=> h'; co_tac.
Ytac0;move=> p1.
ex_middle p2.
move: (factorial_nonzero mB).
have nmB: inc (n -c m) Bnat by hprops.
move: (factorial_nonzero nmB) => z1 z2.
move: (zero_cardinal_product2 p2 z2)=> z3.
move: (zero_cardinal_product2 z3 z1)=> z4.
contradiction.
Qed.

Lemma binom_good: forall n m, inc n Bnat -> inc m Bnat ->
  m <=c n ->
  (binom n m) *c (factorial m) *c (factorial (n -c m)) = (factorial n).
Proof. by move=> n m nB mB h; move: (binom_alt_pr nB mB); Ytac0. Qed.

Lemma binom_pr0: forall n p, inc n Bnat -> inc p Bnat ->
  p <=c n ->
  let num := (factorial n) in
  let den:= (factorial p) *c (factorial (n -c p)) in
  den %|c num & binom n p = num %/c den.
Proof.
move=> n p nB pB h num den; move: (binom_good nB pB h).
rewrite -/num -/den.
have nuB: inc num Bnat by apply BS_factorial =>//.
have deB: inc den Bnat
  by apply BS_mult;apply BS_factorial =>//; hprops.
have dnz: den <> 0c.
   rewrite /den; apply zero_cardinal_product2.
   apply factorial_nonzero=>//.
   apply factorial_nonzero=>//; hprops.
have bN: inc (binom n p) Bnat by apply BS_binom =>//.
rewrite - card_multA card_multC.
move => rel; split.
  rewrite -rel;apply BNdivides_pr1 =>//.
apply BNdivides_pr2 => //.
Qed.

Lemma binom_pr1: forall n p, inc n Bnat -> inc p Bnat ->
  p <=c n ->
  binom n p = (factorial n) %/c ((factorial p) *c (factorial (n -c p))).
Proof.
move=> n p nB pB h; move: (binom_pr0 nB pB h)=> [_]; intuition.
Qed.

Lemma binom_symmetric:forall n p, inc n Bnat -> inc p Bnat ->
  p <=c n -> binom n p = binom n (n -c p).
Proof.
move => n p nB pB h.
move: (sub_le_symmetry h) => aux.
move: (sub_le_symmetry h) => h'.
rewrite (binom_pr1 nB pB h) card_multC (binom_pr1 nB (BS_sub nB pB) h').
rewrite double_sub //; apply le_minus.
Qed.

Lemma binom0: forall n, inc n Bnat -> binom n 0c = 1c.
Proof.
move=> n nB.
case (equal_or_not n 0c).
  by move=> ->; rewrite binom00.
move=> nz; move: (predc_pr nB nz)=> [pB] ->.
rewrite binomSn0 //.
Qed.

Lemma binom1: forall n, inc n Bnat -> binom n 1c = n.
Proof.
rewrite -succ_zero;apply cardinal_c_induction.
  rewrite binom0Sm //; fprops.
move=> p pB r.
have aux:inc 0c Bnat by fprops.
by rewrite (binomSnSm pB aux) r (binom0 pB) (Bsucc_rw pB).
Qed.

Lemma binom2a: forall n, inc n Bnat ->
  2c *c (binom (succ n) 2c) = n *c (succ n).
Proof.
rewrite -{2} succ_one; apply cardinal_c_induction.
  symmetry;rewrite card_multC zero_prod_absorbing.
  rewrite binomSnSm; fprops.
  rewrite binom0Sm; fprops; rewrite binom1;fprops.
  by aw;fprops;rewrite zero_prod_absorbing.
move=> n nB r; move: (BS_succ nB) => sB.
rewrite (binomSnSm sB(inc1_Bnat)).
rewrite cardinal_distrib_prod_sum3 r (binom1 sB).
rewrite - cardinal_distrib_prod_sum2 card_multC.
by congr (_ *c _);rewrite - card_two_pr (Bsucc_rw sB) (Bsucc_rw nB) card_plusA.
Qed.

Lemma binom2: forall n, inc n Bnat ->
  binom (succ n) 2c = (n *c (succ n)) %/c 2c.
Proof.
move=> n nB; move: (BS_succ nB) => sB.
apply (BNdivides_pr2 (BS_mult nB sB) inc2_Bnat (BS_binom sB inc2_Bnat)
  card_two_not_zero).
by rewrite (binom2a nB).
Qed.

Lemma binom_nn : forall n, inc n Bnat -> binom n n = 1c.
Proof.
move=> n nB.
have nn: cardinal_le n n by fprops.
rewrite binom_symmetric// minus_n_nC//; apply binom0 =>//.
Qed.

Lemma binom_pr3: forall n p, inc n Bnat -> inc p Bnat ->
  p <=c n -> binom n p <> 0c.
Proof.
move=> n p nB pB h bz; move: (binom_good nB pB h).
rewrite bz - card_multA card_multC zero_prod_absorbing =>h'.
by elim (factorial_nonzero nB).
Qed.

Lemma finite_sum4_lt: forall a b, inc a Bnat -> inc b Bnat ->
   a <> 0c -> b <c (b +c a).
Proof.
move=> a b aB bB anz.
have az: 0c <c a.
  split; last by intuition.
  apply zero_smallest; fprops.
move: (finite_sum3_lt inc0_Bnat bB aB az).
rewrite (card_plusC a b);aw; fprops.
Qed.

Lemma binom_monotone1: forall k n m,
  inc k Bnat -> inc n Bnat -> inc m Bnat ->
  k <> 0c -> k <=c (succ n) -> n <c m ->
  (binom n k) <c (binom m k).
Proof.
have : forall k n, inc k Bnat -> inc n Bnat ->
  k <> 0c -> k <=c (succ n) -> (binom n k) <c (binom (succ n) k).
  move => k n kB nB r1 r2.
  move: (predc_pr kB r1) => [pB ps].
  rewrite ps binomSnSm; fprops; rewrite -ps.
  apply finite_sum4_lt; try apply BS_binom;fprops.
  apply binom_pr3 =>//.
  rewrite lt_n_succ_le1 //; fprops; rewrite -ps//.
move=> aux k n m kB nB mB r1 r2 r3.
pose r p := (binom n k) <c (binom p k).
have sn: inc (succ n) Bnat by hprops.
have mi: inc m (interval_cc Bnat_order (succ n) m).
  rewrite Bnat_interval_cc_pr1 //; split; fprops.
   rewrite lt_n_succ_le0 //; fprops.
apply (cardinal_c_induction3_v (r:=r) sn mB) =>//.
  rewrite /r; apply aux =>//.
move => n'; rewrite /r.
   rewrite Bnat_interval_co_pr1 //; move => [r4 r5].
have r6: (binom n' k) <c (binom (succ n') k).
  apply aux; fprops.
  move:r5 => [r5 _]; Bnat_tac.
  have r7: cardinal_le k n' by co_tac.
  have cn': is_cardinal n' by move: r7=> [_ []].
  move: (is_le_succ0 cn')=> r8.
  co_tac.
move => r7; co_tac.
Qed.

Lemma binom_monotone2: forall k n m,
  inc k Bnat -> inc n Bnat -> inc m Bnat ->
  k <> 0c -> k <=c (succ n) -> k <=c (succ m) ->
  (n <c m = (binom n k) <c (binom m k)).
Proof.
move=> k n m kB nB mB r1 r2 r3.
apply iff_eq.
   apply binom_monotone1 =>//.
have cn: is_cardinal n by fprops.
have cm: is_cardinal m by fprops.
case (cardinal_le_total_order2 cm cn) => //.
move=> mn [ble bne].
case (equal_or_not n m) => nm.
  by rewrite nm in bne; elim bne.
have lmn : m <c n by split; intuition.
have bgt: (binom m k) <c (binom n k) by apply binom_monotone1 =>//.
co_tac.
Qed.

Lemma number_of_partitions_p3: forall E m p, inc m Bnat -> inc p Bnat ->
  cardinal E = m +c p ->
  cardinal (set_of_partitions (variantLc m p) E) =
  binom (m +c p) m.
Proof.
move=> E m p mB pB cE; move: (number_of_partitions_p2 mB pB cE); simpl.
set x:= cardinal _.
set n := factorial _.
set d := (factorial _) *c _.
move=> [p1 [p2 [p3 [p4 p5]]]].
have r1: m <=c (m +c p).
  by apply sum_increasing3; fprops.
have r2:((m +c p) -c m = p) by rewrite card_plusC; apply card_sub_pr1; fprops.
rewrite card_multC in p2.
rewrite (BNdivides_pr2 p3 p4 p1 p5 p2).
by rewrite (binom_pr1 (BS_plus mB pB) mB r1) r2.
Qed.

Lemma number_of_partitions_p4: forall E n m, inc n Bnat -> inc m Bnat ->
  cardinal E = n ->
  cardinal (set_of_partitions (variantLc m (n -c m)) E) =
  binom n m.
Proof.
move=> E n m nB mB cE.
have cn: is_cardinal n by fprops.
have cm: is_cardinal m by fprops.
case (cardinal_le_total_order2 cm cn) => mn.
  move: (card_sub_pr0 nB mB mn) => [sB eql].
  rewrite -eql in cE.
  by rewrite (number_of_partitions_p3 mB sB cE) eql.
have nmn: ~ m <=c n by move=> aux; co_tac.
rewrite binom_bad // card_sub_wrong => //.
set w := set_of_partitions _ E; case (emptyset_dichot w).
  move=> ->; apply cardinal_emptyset.
move=> [t]; rewrite {1} /w set_of_partitions_rw.
move=> pip; move: (pip_prop0 pip)=> aux.
move:pip=>[df [cV [fgf [duj ue]]]].
have ad: inc TPa (domain (variantLc m 0c)) by bw; fprops.
move: (cV _ ad); bw=> eq1.
have tA:inc TPa (domain t) by rewrite df; bw; fprops.
move: (sub_smaller (aux _ tA) ); rewrite eq1 cE => eq2.
co_tac.
Qed.

Definition subsets_with_p_elements p E:=
  Zo (powerset E)(fun z=> cardinal z =p).

Lemma cardinal_complement1:forall n p E A, inc n Bnat -> inc p Bnat ->
  cardinal E = n -> cardinal A = p -> sub A E ->
  cardinal (complement E A) = n -c p.
Proof.
move=> n p E A nB pB cE cA AE.
move: (cardinal_complement AE) =>cc.
symmetry; rewrite - cE - cA - cc card_plusC cA.
apply card_sub_pr1 =>//.
have ccEA: is_cardinal (cardinal (complement E A)) by fprops.
by apply Bnat_in_sum with (cardinal A) => //; rewrite cc cE.
Qed.

Lemma subsets_with_p_elements_pr: forall n p E, inc n Bnat -> inc p Bnat ->
  cardinal E = n ->
  binom n p = cardinal (subsets_with_p_elements p E).
Proof.
move=> n p E nB pB cEn.
set src := set_of_partitions (variantLc p (card_sub n p)) E.
set trg:= subsets_with_p_elements p E.
have cn: is_cardinal n by fprops.
have cp: is_cardinal p by fprops.
case (cardinal_le_total_order2 cp cn); last first.
  move=> np; rewrite binom_bad//.
  case (emptyset_dichot trg).
     by move=> ->; rewrite cardinal_emptyset.
  move=> [y]; rewrite /trg /subsets_with_p_elements Z_rw; aw.
  move=> [yE cy]; move: (sub_smaller yE); rewrite cEn cy => pn; co_tac.
move=> pn.
move: (card_sub_pr0 nB pB pn) => [sB ps].
move: (cEn); rewrite -ps; move=> cEn'.
rewrite -(number_of_partitions_p3 pB sB cEn') -/src.
clear cEn'; aw.
exists (BL (V TPa) src trg); ee; aw.
apply bl_bijective.
    move=> z; rewrite /src/trg set_of_partitions_rw => pip.
    have ta: inc TPa (domain z).
      move: pip => [p1 _]; rewrite p1; bw; fprops.
    rewrite /subsets_with_p_elements; apply Z_inc.
    move: ((pip_prop0 pip) _ ta); aw.
    move: pip => [_ [cp' _]].
    move: cp'; bw =>cp'; rewrite (cp' _ inc_TPa_two_points); bw.
  have aux: forall w, inc w src -> (V TPb w = complement E (V TPa w)).
    move=> w; rewrite /src set_of_partitions_rw.
    move=>[df [cV [fgf [duj ue]]]].
    have dgw: (domain w= two_points) by rewrite df; bw.
    red in duj; rewrite dgw in duj.
    case (duj _ _ inc_TPa_two_points inc_TPb_two_points) => di1.
      by elim (two_points_distinct).
    set_extens u; rewrite - ue; srw; rewrite dgw.
      move=> uW; split.
         exists TPb; split; ee.
      move=> h; red in di1; empty_tac1 u; apply intersection2_inc => //.
   move=> [[y [ytp uW]] nuW]; move: ytp.
   rewrite two_points_rw; case; move =>eql; rewrite eql in uW =>//.
   move=> u v us vs sWa.
   have sWb: V TPb u = V TPb v by rewrite (aux _ us) (aux _ vs) sWa.
  move: us vs; rewrite /src !set_of_partitions_rw.
  move=>[df [cV [fgf _]]][df' [cV' [fgf' _]]].
  apply fgraph_exten =>//; first by rewrite df'.
  by rewrite df; bw;move=> x;rewrite two_points_rw;case; move=> ->.
move=> y; rewrite /subsets_with_p_elements Z_rw; aw.
move=> [yE cyp].
have aux:partition_fam (partition_with_complement E y) E.
   apply is_partition_with_complement =>//.
exists (variantLc y (complement E y)).
rewrite /src set_of_partitions_rw / partition_with_pi_elements;bw;ee.
move=> i itp; try_lvariant itp.
rewrite (cardinal_complement1 nB pB cEn cyp yE) //.
Qed.

Lemma subsets_with_p_elements_pr0: forall n p, inc n Bnat -> inc p Bnat ->
  binom n p = cardinal (subsets_with_p_elements p n).
Proof.
move=> n p nB pB.
have cn: (cardinal n = n) by apply cardinal_of_cardinal; fprops.
by rewrite (subsets_with_p_elements_pr nB pB cn).
Qed.

Lemma bijective_complement: forall n p E, inc n Bnat -> inc p Bnat ->
  p <=c n -> cardinal E = n ->
  bijection (BL (complement E)
    (subsets_with_p_elements p E)(subsets_with_p_elements (n -c p) E)).
Proof.
move=> n p E nB pB pn cE.
rewrite /subsets_with_p_elements.
apply bl_bijective.
    move=> z; rewrite Z_rw; aw.
    move=> [zE cz]; apply Z_inc.
      aw;apply sub_complement.
    apply cardinal_complement1 =>//.
  move => u v; rewrite Z_rw Z_rw; aw. move=> [h1 _][h2 _].
  rewrite -{2}(double_complement h1) -{2}(double_complement h2).
  by move => ->.
move=> y; rewrite Z_rw; aw; move => [yE cy].
exists (complement E y).
rewrite (double_complement yE).
split =>//; apply Z_inc.
  aw; apply sub_complement.
rewrite (cardinal_complement1 nB (BS_sub nB pB) cE cy yE).
apply double_sub => //.
Qed.

Definition set_of_functions_sum_eq E n:=
  Zo (set_of_functions E (interval_cc_0a n))
  (fun z=> (cardinal_sum (P z)) = n).

Definition set_of_graph_sum_eq F n:=
  fun_image (set_of_functions_sum_eq F n) graph.

Definition set_of_functions_sum_le E n:=
  Zo (set_of_functions E (interval_cc_0a n))
  (fun z=> (cardinal_sum (P z)) <=c n).

Definition set_of_graph_sum_le F n:=
  fun_image (set_of_functions_sum_le F n) graph.

Lemma setof_suml_aux: forall F n f, inc n Bnat ->
  inc f (set_of_graph_sum_le F n) =
  (fgraph f & domain f = F & (cardinal_sum f) <=c n &
    (forall i, inc i (domain f) -> is_cardinal (V i f))).
Proof.
move=> F n f nB; rewrite /set_of_graph_sum_le /set_of_functions_sum_le.
apply iff_eq; aw.
  move=> [z]; rewrite Z_rw; aw; move=> [[[ fz [sz tz]] cs] gz].
  rewrite - gz; ee.
    rewrite -sz; aw.
  aw; move=> i iz; move: (inc_W_target fz iz); rewrite tz => WB.
  move: (sub_interval_Bnat WB); move=> h; fprops.
move=> [fgf [df [sf alc]]].
have tf: forall i, inc i F -> inc (V i f) (interval_Bnat 0c n).
  rewrite -df;move=> i iF.
  set (j := singleton i).
  have sjd: sub j (domain f) by move=> t; rewrite singleton_rw; move=> ->.
  move: (alc _ iF) => aux.
  move: (sum_increasing1 fgf alc sjd).
  have vr: V i f = V i (restr f j) by bw; rewrite /j; fprops.
  rewrite vr in aux.
  have dr: (domain (restr f j) = singleton i) by rewrite restr_domain1 //.
  rewrite (trivial_cardinal_sum1 dr aux).
  rewrite -vr; move=> vle0.
  have vle: cardinal_le (V i f) n by co_tac.
  rewrite interval_Bnat_pr0 =>//.
set z:=corresp F (interval_cc_0a n) f.
exists z; rewrite /z;ee; aw.
apply Z_inc;aw; ee.
  apply is_function_pr;ee.
  move=> t; rewrite frange_inc_rw //.
  move=> [x [xs ]] ->; apply tf; ue.
rewrite /corresp; aw.
Qed.

Lemma setof_sume_aux: forall F n f, inc n Bnat ->
  inc f (set_of_graph_sum_eq F n) =
  (fgraph f & domain f = F & cardinal_sum f = n &
    (forall i, inc i (domain f) -> is_cardinal (V i f))).
Proof.
move=> F n f nB; rewrite /set_of_graph_sum_eq /set_of_functions_sum_eq.
apply iff_eq; aw.
  move=> [z]; rewrite Z_rw; aw; move=> [[[ fz [sz tz]] cs] gz].
  rewrite - gz; ee.
    rewrite -sz; aw.
  aw; move=> i iz; move: (inc_W_target fz iz); rewrite tz => WB.
  move: (sub_interval_Bnat WB); move=> h; fprops.
move=> [fgf [df [sf alc]]].
have : inc f (set_of_graph_sum_le F n).
  rewrite setof_suml_aux; ee; rewrite sf; fprops.
rewrite /set_of_graph_sum_le /set_of_functions_sum_le; aw.
move=> [z []]; rewrite Z_rw; move=> [zs cle] gz.
by exists z; ee; rewrite Z_rw; ee; rewrite -/(graph z) gz.
Qed.

Lemma sum_of_gen_binom: forall E F n, inc n Bnat -> cardinal E = n ->
  cardinal_sum (L (set_of_graph_sum_eq F n)
              (fun p => cardinal (set_of_partitions p E)))
  = (cardinal F) ^c n.
Proof.
move=> E F n cE cEn.
have ->: ((cardinal F) ^c n = F ^c E).
  rewrite - cEn;apply card_pow_pr; fprops.
pose g f i:= (inv_image_by_fun f (singleton i)).
set f1:=fun f => (L F (fun i=> cardinal (g f i))).
set f2:= fun f => L F (g f).
have p1: forall f, inc f (set_of_functions E F) ->
     (inc (f1 f) (set_of_graph_sum_eq F n)
      & inc (f2 f) (set_of_partitions (f1 f) E)).
  move=> f; aw; move=> [ff [sf tf]].
  have sfa: (source f = unionb (L (target f) (g f))).
    rewrite /g;set_extens t; srw; bw => // ts.
      move: (inc_W_target ff ts)=> Wt.
      by ex_tac; bw;rewrite /inv_image_by_fun; aw; ex_tac; apply W_pr3.
    move: ts=> [y [ytg]]; bw; rewrite /inv_image_by_fun; aw.
    move=> [u]; aw; move=> [eq Jg].
    rewrite eq in Jg; apply (inc_pr1graph_source ff Jg).
  have md: mutually_disjoint (L (target f) (g f)).
    apply mutually_disjoint_prop2.
    move=> i j y itf jtf; rewrite /g/inv_image_by_fun; aw.
    move=> [u [ui ug]] [v [vi vg]]; move: ui vi; aw; move=> <- <-.
    by rewrite -(W_pr ff ug) -(W_pr ff vg).
  rewrite /f1 /f2;split.
    rewrite setof_sume_aux //; bw;ee; last first.
      move=> i iF; bw; fprops.
    change (cardinal_sum (L F (fun i : Set => cardinal (g f i))) = n).
    rewrite - cEn -sf sfa.
    rewrite /cardinal_sum; aw; rewrite /disjoint_union.
    apply equipotent_disjoint_union; gprops.
    - rewrite /disjoint_union_fam; gprops.
    - rewrite /disjoint_union_fam tf; bw.
    - rewrite /disjoint_union_fam -tf; bw; move=> i itf; bw.
      eqtrans (cardinal (g f i)); eqsym; fprops.
    - apply disjoint_union_disjoint; gprops.
  rewrite set_of_partitions_rw; red; bw; ee.
    move => i iF; bw.
  red; aw;ee; rewrite -tf /BL; aw.
  by rewrite - sfa.
set G:= L (set_of_graph_sum_eq F n)
    (fun p => (Zo (set_of_functions E F) (fun f=> f1 f = p))).
have g1: unionb G = (set_of_functions E F).
  rewrite /G;set_extens f; srw; bw.
    move=> [y [ys]]; bw; rewrite Z_rw; intuition.
  by move=> fsf;move: (p1 _ fsf) => [fa fb]; ex_tac;bw; apply Z_inc.
have g2: mutually_disjoint G.
  rewrite /G;apply mutually_disjoint_prop2.
  move=> i j y ip jp; rewrite Z_rw Z_rw.
  rewrite /f1/set_of_partitions /partition_with_pi_elements.
  by move=> [_ r1][_ r2]; rewrite -r1 -r2.
symmetry; rewrite /card_pow; aw.
rewrite -g1 /cardinal_sum /disjoint_union.
aw;rewrite /G;apply equipotent_disjoint_union; gprops;
   last by apply disjoint_union_disjoint; gprops.
- rewrite /disjoint_union_fam; gprops.
- rewrite /disjoint_union_fam; bw.
- rewrite /disjoint_union_fam;bw; move=> i id; bw.
eqtrans (cardinal (set_of_partitions i E)).
eqtrans (set_of_partitions i E).
rewrite /set_of_partitions / partition_with_pi_elements.
set X:= Zo _ _.
exists (BL f2 X (set_of_partitions i E)); aw;ee.
apply bl_bijective.
    rewrite /W; move=> t; rewrite Z_rw.
    move=> [ts] <-; move : (p1 _ ts); intuition.
  move=> u v ;rewrite /X Z_rw Z_rw.
  aw; move=> [[fu [su tu]] f1u] [[fv [sv tv]] f1v].
  rewrite /f2 /g => Weq.
  apply function_exten =>//; try ue.
  move=> x xsu.
  have WF: inc (W x u) F by rewrite -tu; apply inc_W_target=>//.
  move: (f_equal (V (W x u)) Weq); bw => ieq.
  have : (inc x (inv_image_by_fun u (singleton (W x u)))).
    rewrite /inv_image_by_fun; aw; ex_tac; apply W_pr3 =>//.
  rewrite ieq; rewrite /inv_image_by_fun; aw.
  by move=> [w]; aw; move=> [wW Jg]; rewrite -wW (W_pr fv Jg).
move=> y; rewrite /X /f1 /f2 set_of_partitions_rw => ys.
move: (pip_prop0 ys)=> pip0.
move: ys=>[df [cV [fgf [duj ue]]]].
have doi: domain i = F.
  by move: id; rewrite(@setof_sume_aux F n i cE); move=> [fgi [di _]].
rewrite doi in cV.
set xf := fun x => choose (fun j => inc j F & inc x (V j y)).
have xfp: forall x, inc x E -> (inc (xf x) F & inc x (V (xf x) y)).
  move=> x xE; apply choose_pr.
  by move: xE;rewrite -ue; srw; aw; rewrite df doi.
have ta: transf_axioms xf E F by move=> t tE; move: (xfp _ tE); intuition.
set x:= BL xf E F.
have fx: is_function x by apply bl_function.
have xsf: inc x (set_of_functions E F).
   by rewrite /x;aw; ee.
have xp1: forall t, inc t F -> g x t = V t y.
  move=> t tF; rewrite /g.
  rewrite /inv_image_by_fun;set_extens u; aw.
    move=> [v]; aw; move=> [vt Jg]; rewrite vt in Jg.
    move: (W_pr fx Jg).
    move: (inc_pr1graph_source fx Jg); rewrite /x bl_source => uE.
    by aw; move=> <-; move: (xfp _ uE) => [_].
  move=> uW; ex_tac.
  rewrite df doi in pip0; move: (pip0 _ tF)=> vE; move: (vE _ uW) => uE.
  move: (xfp _ uE) => [vF tv].
  move: duj; rewrite /mutually_disjoint;aw; rewrite df doi.
  move=> aux; case (aux _ _ tF vF) => a.
    have: t = W u x by rewrite /x; aw.
    move=> ->; apply W_pr3 =>//; rewrite /x; aw; ue.
  by red in a; empty_tac1 v; apply intersection2_inc.
have xp2:L F (g x) = y.
  symmetry; apply fgraph_exten =>//; bw; gprops.
    rewrite df //.
  rewrite df doi; move => u uF; bw; rewrite - xp1 //.
exists x; split=>//.
apply Z_inc =>//.
apply fgraph_exten; gprops; bw.
    by move: id; rewrite(@setof_sume_aux F n i cE); move=> [fgi _].
   by symmetry.
move=> u uF; bw; rewrite xp1// cV //.
Qed.

Lemma sum_of_gen_binom0: forall E n a,
  inc n Bnat -> cardinal E = n -> finite_int_fam a ->
  (cardinal_sum a) ^c n =
     cardinal_sum (L (set_of_graph_sum_eq (domain a) n)
       (fun p =>
         (cardinal (set_of_partitions p E)) *c
            (cardinal_prod (L (domain a) (fun i=> ((V i a) ^c (V i p))))))).
Proof.
move=> E n a nB cEn fifa.
set F:=cardinal_sum a; set I := domain a.
move: (finite_sum_finite fifa) => sB.
have cF: cardinal_sum a = cardinal F .
  by symmetry; apply cardinal_of_cardinal;apply Bnat_is_cardinal.
move: (number_of_partitions1 fifa cF) => [pF].
rewrite set_of_partitions_rw => pipF.
have ->: (F ^c n = F ^c E).
  rewrite - cEn;apply card_pow_pr; fprops.
pose g f i := (inv_image_by_fun f (V i pF)).
pose f1 f := (L I (fun i=> cardinal (g f i))).
pose f2 f := L I (g f).
have g_prop: forall f i, inc f (set_of_functions E F) -> inc i I ->
    forall x , inc x (g f i) = (inc x E & inc (W x f) (V i pF)).
  move=> f i; aw; move=> [ff [sf tf]] iI x.
  rewrite /g /inv_image_by_fun; aw.
  rewrite -sf;apply iff_eq.
    move=> [u [uV Jg]]; split; [ graph_tac | rewrite (W_pr ff Jg) //].
  move=> [xE Wv]; ex_tac; apply W_pr3 =>//.
move: (pip_prop0 pipF) => prop0F.
move:pipF=>[da [cVa [fga [duja uea]]]].
have sfa: forall f, inc f (set_of_functions E F) -> unionb (L I (g f)) = E.
  move=> f fs; move: (fs); aw; move=> [ff [sf tf]].
  symmetry; rewrite -sf;set_extens t; srw; bw => // ts.
    move: (inc_W_target ff ts).
    rewrite tf -uea; srw; rewrite da; move=> [y [yda Vy]].
    exists y; ee; bw; rewrite (g_prop _ _ fs yda); split=>//; ue.
  move: ts => [y [yI]]; bw; rewrite (g_prop _ _ fs yI) sf; intuition.
have sfb: forall f, inc f (set_of_functions E F) ->
      mutually_disjoint (L I (g f)).
    move=> f fs;apply mutually_disjoint_prop2.
    move=> i j y iI jI; rewrite (g_prop _ _ fs iI) (g_prop _ _ fs jI).
    move=> [yE W1][_ W2].
    red in duja;move: duja;rewrite da; move=> h; case (h _ _ iI jI)=>//.
    rewrite /disjoint=> h'; empty_tac1 (W y f); apply intersection2_inc =>//.
have sfc:forall f, inc f (set_of_functions E F) ->
     inc (f2 f) (set_of_partitions (f1 f) E).
  move=> f fs; rewrite /f1 /f2 set_of_partitions_rw; red; bw;ee.
    move => i iF; bw.
  red; aw;ee.
have sfd: forall f, inc f (set_of_functions E F) ->
    inc (f1 f) (set_of_graph_sum_eq I n).
  move=> f fs; rewrite /f1 setof_sume_aux //; bw; ee.
  rewrite - cEn -(sfa _ fs).
    rewrite /cardinal_sum; aw; rewrite /disjoint_union.
    apply equipotent_disjoint_union; bw;gprops.
    - rewrite /disjoint_union_fam; gprops.
    - rewrite /disjoint_union_fam; bw.
    - rewrite /disjoint_union_fam; bw; move=> i itf; bw.
      eqtrans (cardinal (g f i)); eqsym; fprops.
    - apply disjoint_union_disjoint; gprops.
  move=> i iI; bw; fprops.
set G:= L (set_of_graph_sum_eq I n)
    (fun p => (Zo (set_of_functions E F) (fun f=> f1 f = p))).
have g1: unionb G = (set_of_functions E F).
  rewrite /G;set_extens f; srw; bw.
    move=> [y [ys]]; bw; rewrite Z_rw; intuition.
  by move=> fsf;move: (sfd _ fsf) => fa; ex_tac; bw; apply Z_inc.
have g2: mutually_disjoint G.
  rewrite /G;apply mutually_disjoint_prop2.
  move=> i j y ip jp; rewrite Z_rw Z_rw.
  rewrite /f1/set_of_partitions /partition_with_pi_elements.
  by move=> [_ r1][_ r2]; rewrite -r1 -r2.
symmetry; rewrite /card_pow; aw.
rewrite -g1 /cardinal_sum /disjoint_union.
aw;rewrite /G;apply equipotent_disjoint_union; gprops;
   last by apply disjoint_union_disjoint; gprops.
- rewrite /disjoint_union_fam; gprops.
- rewrite /disjoint_union_fam; bw.
- rewrite /disjoint_union_fam;bw; move=> p pd; bw.
set X:= (set_of_partitions p E).
set Y := productb (L I (fun i=> (set_of_functions (V i p) (V i pF)))).
set aux:= card_mult _ _.
eqtrans aux; first by eqsym; fprops.
eqtrans (product X Y).
  rewrite /aux card_mult_pr1/cardinal_prod; set t:= product _ _; eqtrans t.
  rewrite /t; clear t; apply equipotent_product; fprops.
  rewrite /cardinal_prod.
  set t:= productb _ ; eqtrans t.
  rewrite /t; clear t; apply equipotent_productb; bw;gprops.
  move=> i iI; bw; fprops.
  eqtrans (cardinal (set_of_functions (V i p) (V i pF))).
  change (equipotent (card_pow (V i a) (V i p)) (card_pow (V i pF) (V i p))).
  have ->: (card_pow (V i a) (V i p) = card_pow (V i pF) (V i p)).
     apply card_pow_pr; fprops; rewrite - cVa //; fprops.
  fprops.
clear aux.
set (Z:= (Zo (set_of_functions E F) (fun f : Set => f1 f = p))).
have phi1: forall f, inc f Z -> inc (f2 f) X.
   move=> f; rewrite /Z Z_rw; move=> [fs f1f].
  rewrite /X -f1f; apply sfc =>//.
set WI:= fun f i => equipotent_ex (V i p) (g f i).
have W1: forall f i, inc f Z -> inc i I ->
    (bijection (WI f i) & source (WI f i) = (V i p) & target (WI f i)= g f i).
  move=> f i fZ iI; rewrite /WI; apply equipotent_ex_pr.
  move: fZ; rewrite /Z Z_rw /f1; move=> [fa fb]; rewrite -fb; bw; fprops.
have W2: forall f f' i, inc f Z -> inc f' Z-> inc i I ->
    f2 f = f2 f' -> WI f i = WI f' i.
  move=> f f' i fZ f'Z iI sf1; rewrite /WI.
  by move: (f_equal (V i) sf1); rewrite /f2; bw; move=> ->.
set WJ:= fun f i => (restriction2 f (g f i) (V i pF)) \co (WI f i).
have W3: forall f i, inc f Z -> inc i I ->
  (restriction2_axioms f (g f i) (V i pF) &
   is_function (restriction2 f (g f i) (V i pF)) &
   (restriction2 f (g f i) (V i pF)) \coP (WI f i)).
  move=> f i fZ iI.
  move: (W1 _ _ fZ iI) => [bW [sW tW]].
  move: fZ; rewrite /Z Z_rw; move=> [feF f1f].
  move: (g_prop _ _ feF iI) => g1p.
  move: feF;aw; move=> [ff [sf tf]].
  have sgi: sub (g f i) (source f).
     by rewrite sf; move=> t; rewrite g1p; intuition.
  have ra:restriction2_axioms f (g f i) (V i pF).
    red; ee.
      rewrite tf; apply prop0F; ue.
    by move=> t; aw; move=> [u]; rewrite g1p; move=> [[uE pr]] <-.
  have fr:is_function (restriction2 f (g f i) (V i pF)).
    apply restriction2_function =>//.
    by ee; red; ee; [ fct_tac | rewrite /restriction2;aw].
set f3:= fun f => (L I (WJ f)).
have phi2: forall f, inc f Z -> inc (f3 f) Y.
  move=> f fZ; rewrite /Y /f3 productb_rw//; bw;ee.
  move=> i iV; bw.
  move: (W3 _ _ fZ iV)(W1 _ _ fZ iV) => [p1 [p2 p3]][p4 [p5 p6]].
  rewrite /WJ; aw;ee; [ fct_tac | rewrite /restriction2; aw].
eqsym.
exists (BL (fun f =>(J (f2 f) (f3 f))) Z (product X Y)).
aw;ee; apply bl_bijective.
   move=> f fZ; aw;ee.
  move=> u v uZ vZ eq.
  move: (pr1_def eq)(pr2_def eq)=> sf2 sf3.
  move: (uZ); rewrite /Z Z_rw; move => [usf _].
  move: (uZ)(vZ); rewrite /Z Z_rw Z_rw; aw.
  move=> [[fu [su tu]] f1u] [[fv [sv tv]] f1v].
  apply function_exten =>//; try ue.
  rewrite su -(sfa _ usf); move=> x; srw; bw.
  move=> [i [iI]]; bw => xgui.
  move: sf3; rewrite /f3 => eq1.
  move: (f_equal (V i) eq1); bw; rewrite /WJ => eq2.
  move:(W2 _ _ _ uZ vZ iI sf2) => eq3.
  move: (W3 _ _ uZ iI)(W1 _ _ uZ iI) => [p1 [p2 p3]][p4 [p5 p6]].
  move: (W3 _ _ vZ iI)(W1 _ _ vZ iI) => [p7 [p8 p9]][p10 [p11 p12]].
  rewrite -p6 in xgui.
  move: p4 => [_ sW]; move: (surjective_pr2 sW xgui) => [y [ys vy]].
  move: (f_equal (W y) eq2); aw; try ue.
  rewrite (restriction2_W p1); last by rewrite vy -p6.
  rewrite (restriction2_W p7).
    rewrite - eq3 vy //.
  rewrite -p12; apply inc_W_target; last by ue.
  by fct_tac; rewrite p11 -p5.
move=> y; aw; move=> [py [Py Qy]].
move: (Py); rewrite /X set_of_partitions_rw.
move=> [dy [cVy pfay]]; move: (pfay) =>[fgy [dujy uey]].
move: (Qy); rewrite /Y Z_rw; bw; move=> [_ [fgQ [dQ Qv]]].
move: (pd); rewrite (setof_sume_aux _ _ nB); move => [fgp [dp [csp cVip]]].
set WK:= fun i => equipotent_ex (V i p) (V i (P y)).
have W4: forall i, inc i I ->
  (bijection (WK i) & source (WK i) = V i p & target (WK i) = (V i (P y))).
  move=> i iI; rewrite /WK; apply equipotent_ex_pr; rewrite - cVy; fprops; ue.
pose hb i := (canonical_injection (V i pF) F) \co
  ((V i (Q y)) \co (inverse_fun (WK i))).
have Hv:forall i, inc i I ->
  ((V i (Q y)) \coP (inverse_fun (WK i)) &
   is_function ( (V i (Q y)) \co (inverse_fun (WK i)))
   & (canonical_injection (V i pF) F) \coP
  ((V i (Q y)) \co (inverse_fun (WK i)))).
  move=> i iI.
  move: (W4 _ iI)(Qv _ iI); bw;aw; move => [bW [sw tW]][fV [sV tV]].
  have p1: (V i (Q y)) \coP (inverse_fun (WK i)).
    red; ee; [ by apply bijective_inv_function | aw; ue].
  have p2:is_function ((V i (Q y)) \co (inverse_fun (WK i))) by fct_tac.
  ee; red; ee.
    apply ci_function; apply prop0F; ue.
  by rewrite /canonical_injection; aw.
have Hw:forall i, inc i I -> function_prop (hb i) (V i (P y)) F.
  move=> i iI.
  move: (Hv _ iI) => [p1 [p2 p3]]; rewrite /hb;red.
   rewrite compose_source compose_source compose_target.
  split; first by fct_tac.
  split.
    rewrite inverse_source; by move: (W4 _ iI); move => [bW [sw tW]].
  rewrite /canonical_injection; aw.
rewrite -dp -dy in Hw.
move: (extension_partition pfay Hw) => [[w [[fw [sw tw]] alaw]] _].
rewrite dy dp in alaw.
have wsf:inc w (set_of_functions E F) by aw.
have gp1: forall j x, inc j I -> inc x (V j (P y)) ->
    inc (W x w) (V j pF).
  move=> j x jI xV.
  move: (alaw _ jI) =>[p1 [p2 p3]]; rewrite (p3 _ xV).
  move: (Hv _ jI) => [p4 [p5 p6]].
  move: (W4 _ jI) (Qv _ jI); bw;aw; move => [bWj [sWj tWj]][fVj [sVj tVj]].
  rewrite /hb.
  set aux := (V j (Q y)) \co _.
  have xsa: inc x (source aux) by rewrite /aux; aw; rewrite tWj.
  have xta: inc (W x aux) (V j pF).
       have <-:(target aux = V j pF) by rewrite /aux; aw.
       by apply inc_W_target.
    rewrite compose_W // ci_W //.
    apply prop0F; ue.
have gp2: forall i, inc i I -> g w i = V i (P y).
  move => i iI.
  set_extens x; rewrite (g_prop _ _ wsf iI).
  move=> [xE Wp].
  move: xE; rewrite - uey; srw; move=> [j [jI xVj]].
  rewrite dy dp in jI.
  move: (gp1 _ _ jI xVj) => Wp'.
  red in duja; rewrite da in duja; case (duja _ _ iI jI) => h.
    by rewrite h.
  by red in h; empty_tac1 (W x w); apply intersection2_inc.
  move: (alaw _ iI)=> [p1 [p2 p3]].
  split; first by rewrite - sw; apply p1.
  by apply gp1.
have wZ:inc w Z.
  rewrite /Z Z_rw;ee.
  apply fgraph_exten =>//;rewrite /f1; bw; gprops; try ue.
  move=> x xI; bw; rewrite gp2 //; apply cVy; ue.
exists w; ee.
have ->: (f2 w = P y).
  apply fgraph_exten =>//;rewrite /f2; bw; gprops; try ue.
  move=> x xI; bw; rewrite gp2 =>//.
suff: (f3 w = Q y) by move=> ->; aw.
apply fgraph_exten =>//;rewrite /f3; bw; gprops; try ue.
move => i iI; bw.
move: (W4 _ iI)(Qv _ iI); bw;aw; move => [bW [sW tW]][fV [sV tV]].
have ww: (WI w i = WK i) by rewrite /WI /WK gp2 //.
move: (W3 _ _ wZ iI); rewrite /WJ ww (gp2 _ iI); move => [ra [rd cr]].
apply function_exten; rewrite ?compose_source.
  - apply compose_function =>//; try (aw; ue).
  - exact fV.
  - ue.
  - aw; rewrite /restriction2; aw; ue.
have fW: is_function (WK i) by fct_tac.
move=> x xs.
have p1:inc (W x (WK i)) (target (WK i)) by apply inc_W_target.
move: (p1); rewrite tW => p11.
rewrite compose_W //restriction2_W //.
move: (alaw _ iI)=> [_ [_ p3]]; rewrite -tW in p3;rewrite (p3 _ p1).
rewrite /hb.
have p2: (V i (Q y)) \coP (inverse_fun (WK i)).
  red; ee; [ by apply bijective_inv_function | aw; ue].
have p0:is_function ((V i (Q y)) \co (inverse_fun (WK i))) by fct_tac.
rewrite da in prop0F; move: (@prop0F i iI)=> pf0.
have p4: (canonical_injection (V i pF) F) \coP
   ( (V i (Q y)) \co (inverse_fun (WK i))).
  red; ee.
  by rewrite /canonical_injection; aw.
aw.
have <-: x = W (W x (WK i)) (inverse_fun (WK i)) by apply W_inverse2 =>//.
rewrite ci_W // -tV //.
apply inc_W_target =>//.
by rewrite sV -sW.
Qed.

Lemma sum_of_gen_binom2: forall n, inc n Bnat ->
  cardinal_sum (L (interval_cc_0a n) (binom n)) = 2c ^c n.
Proof.
move=> n nB.
have cE: cardinal n = n by apply cardinal_of_cardinal ; fprops.
move: (sum_of_gen_binom two_points nB cE).
have ->: (cardinal two_points = card_two).
  rewrite -two_points_pr2; apply cardinal_doubleton; fprops.
move => <-.
set f:= fun m => (variantLc m (card_sub n m)).
have p1: forall m, inc m (interval_cc_0a n) ->
   inc (f m) (set_of_graph_sum_eq two_points n).
  move=> m; srw => mn.
  have mB: inc m Bnat by Bnat_tac.
  rewrite /f setof_sume_aux;ee; first by bw.
    apply (card_sub_pr mn).
  move: (BS_sub nB mB)=> sB.
  bw;move=> i itp; try_lvariant itp; fprops.
set g := BL f (interval_cc_0a n)(set_of_graph_sum_eq two_points n).
symmetry.
rewrite (cardinal_sum_commutative2 (I:=interval_cc_0a n) (f:= f));
  first last.
- gprops.
- bw; move=> y; rewrite (@setof_sume_aux two_points n y nB).
  move => [fgy [dy [sy alc]]].
  rewrite dy in alc.
  move: (alc _ inc_TPa_two_points) (alc _ inc_TPb_two_points).
  set a := V TPa y; set b := V TPb y.
  move=> ca cb.
  have abn: a +c b = n.
     rewrite -sy; apply card_plus_pr.
     rewrite /a /b; exists TPa; exists TPb; rewrite dy;ee.
     rewrite two_points_pr2 //.
  move: (nB); rewrite -abn => nB'.
  move: (Bnat_in_sum2 ca nB')(Bnat_in_sum cb nB') => aB bB.
  move: (sum_increasing3 ca cb); rewrite abn => an.
  exists a; rewrite /f; split.
    by rewrite interval_Bnat_pr0.
  apply fgraph_exten; bw => //.
    fprops.
  move=> x xtp; try_lvariant xtp.
  rewrite -/b -abn card_plusC card_sub_pr1 //.
- rewrite /f;move=> u v ui vi s.
  move: (f_equal (V TPa) s); bw.
- bw.
- apply f_equal; apply L_exten1 =>//.
move=> x xI; bw; last by apply p1.
rewrite number_of_partitions_p4 //.
apply (@sub_interval_Bnat 0c n _ xI).
Qed.

Lemma sum_of_binomial: forall n, inc n Bnat ->
  cardinal_sum (L (interval_cc_0a n) (binom n)) = 2c ^c n.
Proof.
move=> n nB.
rewrite - card_powerset.
set (idx:= interval_cc_0a n).
set (X:= L idx (fun p=>subsets_with_p_elements p n)).
set (Y:= disjoint_union_fam X).
have fgX: fgraph X by rewrite/ X; gprops.
have fgY: fgraph Y by rewrite /Y/disjoint_union_fam; gprops.
have dx: domain X = domain Y by rewrite /Y/disjoint_union_fam; bw.
have eqv: forall i, inc i (domain X) -> equipotent (V i X) (V i Y).
  move=> i idX; rewrite /Y/disjoint_union_fam; bw;fprops.
have mdX: mutually_disjoint X.
   apply mutually_disjoint_prop2.
   move=> i j y id jd;rewrite / subsets_with_p_elements.
   by rewrite Z_rw Z_rw; move=> [_ yi][_ yj]; rewrite -yi -yj.
have mdY: mutually_disjoint Y by rewrite /Y; fprops.
move:(equipotent_disjoint_union fgX fgY dx eqv mdX mdY).
have ->:unionb X = powerset n.
   rewrite /X; set_extens t; srw; bw.
     by move=> [y [ydx]]; bw; rewrite /subsets_with_p_elements Z_rw; intuition.
  aw;move=> tp; exists (cardinal t).
  have cxi: (inc (cardinal t) idx).
   rewrite /idx interval_Bnat_pr0 //.
   move: (sub_smaller tp) => ltn.
   by have <-: (cardinal n = n) by apply cardinal_of_cardinal;fprops.
  split =>//; bw.
  rewrite /subsets_with_p_elements; apply Z_inc =>//; aw.
rewrite - cardinal_equipotent; move=> ->.
rewrite /Y.
change (cardinal_sum (L idx (binom n)) = cardinal_sum X).
apply cardinal_sum_pr3; gprops; rewrite /X;bw.
move=> i ix; bw.
move: (@sub_interval_Bnat 0c n _ ix) => iB.
rewrite -subsets_with_p_elements_pr0 //.
have aux: inc (binom n i) Bnat by apply BS_binom.
apply cardinal_of_cardinal;fprops.
Qed.

Lemma sum_of_binomal2: forall a b n,
  is_cardinal a -> is_cardinal b -> inc n Bnat ->
  cardinal_sum (L (interval_cc_0a n)
         (fun p => (binom n p) *c (a ^c p) *c (b ^c (n -c p))))
  = (a +c b) ^c n.
Proof.
move=> a b n ca cb nB.
move: n nB; apply cardinal_c_induction.
  rewrite power_x_0.
  set f := L _ _.
  have zb: inc 0c Bnat by fprops.
  have zi: inc 0c (interval_cc_0a 0c) by srw; fprops.
  have df: domain f = singleton 0c.
    rewrite /f; bw.
    set_extens t; aw; [srw; apply zero_smallest2 |by move=>-> ].
  have xo: V 0c f = card_one.
    rewrite /f; bw.
    rewrite minus_n_nC // power_x_0 // power_x_0 // binom00; aw; fprops.
  have cv: is_cardinal (V 0c f) by rewrite xo; fprops.
  by rewrite (trivial_cardinal_sum1 df cv).
move=> n nB.
set fn := L _ _.
move=> hrec.
set X:=((a +c b) ^c n).
set I := (interval_cc_0a n).
have IB:sub I Bnat by apply sub_interval_Bnat.
pose q0 i := (a ^c i) *c (b ^c (n -c i)).
pose q1 i := (a ^c (succ i)) *c (b ^c ((succ n) -c (succ i))).
pose f1 i := (binom n (succ i)) *c (q1 i).
pose f2 i := (binom n i) *c (q1 i).
pose g1 i := (binom n i) *c (a ^c i) *c (b ^c ((succ n) -c i)).
have P1: cardinal_sum (L I f2) = card_mult a X.
  pose g i:= a *c (binom n i) *c (q0 i).
  have ->: L I f2 = L I g.
    apply L_exten1; first by [].
    move=> x xi; rewrite /f2 /g (card_multC a _) - card_multA.
    congr (_ *c _); rewrite card_multC /q1 /q0.
    rewrite (card_sub_pr6 nB (IB _ xi)).
    rewrite (Bsucc_rw (IB _ xi)) power_of_sum2 power_x_1 //.
    rewrite - !card_multA; congr (_ *c _); apply card_multC.
  symmetry;rewrite /X -hrec.
  have fgn: fgraph fn by rewrite /fn; gprops.
  rewrite (cardinal_distrib_prod2_sum ca fgn).
  rewrite /fn L_domain -/I; apply f_equal; apply L_exten1; first by done.
  move=> x xI; bw.
  rewrite card_multA card_multA - card_multA //.
have P2: cardinal_sum (L I g1) = card_mult b X.
  set (g:= fun i=> b *c ((binom n i) *c (q0 i))).
  have ->: L I g1 = L I g.
    apply L_exten1; first by [].
    move=> x xi; rewrite /g1/g (card_multC b _) - ! card_multA.
    congr (_ *c (_ *c _)).
    move: (IB _ xi) => xB.
    move: (Bsucc_rw (BS_sub nB xB)) => aux.
    have ->: ((succ n) -c x) = succ (n -c x).
      apply card_sub_pr2; hprops.
      rewrite /I (interval_Bnat_pr0 _ nB) in xi.
      rewrite aux card_plusC card_plusA card_sub_pr //.
      by rewrite (Bsucc_rw nB).
    rewrite aux power_of_sum2 power_x_1; fprops.
  symmetry;rewrite /X -hrec.
  have fgn: fgraph fn by rewrite /fn; gprops.
  rewrite (cardinal_distrib_prod2_sum cb fgn).
  rewrite /fn L_domain -/I; apply f_equal; apply L_exten1; first by [].
  move=> x xI; rewrite (L_V_rw _ xI) /g /q0.
  congr (_ *c _); symmetry; apply card_multA.
clear hrec fn.
rewrite fct_sum_rec1//.
have sn: inc (succ n) Bnat by hprops.
have a1: is_cardinal (card_pow b (succ n)) by rewrite /card_pow; fprops.
rewrite (binomSn0 nB) power_x_0 (minus_n_0C sn).
rewrite (one_unit_prodl cardinal1) (one_unit_prodl a1).
set ffn:= fun i => _.
move: (sum_of_sums f1 f2 I).
have ->:L I (fun i => (f1 i) +c (f2 i)) = (L I ffn).
  apply L_exten1; first by [].
  move=> i iI.
  have iB: inc i Bnat by apply IB.
  rewrite /ffn -/(q1 i)/f1 /f2 (binomSnSm nB iB).
  rewrite - card_multA cardinal_distrib_prod_sum2 //.
have g1z: (g1 0c) = (card_pow b (succ n)).
   rewrite /g1 power_x_0 (minus_n_0C sn) (binom0 nB).
   set (w:=(b ^c (succ n))).
   have cw: is_cardinal w by rewrite /w/card_pow; fprops.
   by rewrite (one_unit_prodl cardinal1)(one_unit_prodl cw).
rewrite -g1z.
move=> <-.
set (P11:= cardinal_sum (L I f2)).
have ->: P11 = a *c X by apply P1.
clear P1 P11.
rewrite (card_plusC _ (a *c X)) - card_plusA.
have cz:is_cardinal (g1 0c) by rewrite g1z /card_pow; fprops.
move: (fct_sum_rec1 g1 nB).
have P3: cardinal_sum (L (interval_cc_0a (succ n)) g1)=
  card_mult b X.
  rewrite -P2 /I.
  have gp: (forall a, inc a Bnat -> is_cardinal (g1 a)).
  move=> u uB; rewrite /g1; apply card_mult_is_cardinal.
  move: (induction_on_sum g1 sn); simpl.
  rewrite (interval_co_cc sn) (interval_co_cc nB).
  move=> <-.
  move: (is_lt_succ nB) => aux.
  have ->: (g1 (succ n) = 0c).
    rewrite /g1 (binom_bad nB sn aux) - card_multA card_multC.
    apply zero_prod_absorbing.
  rewrite zero_unit_sumr; [done | rewrite /cardinal_sum; fprops].
have p1: L I (fun i : Set => g1 (succ i)) = (L I f1).
  apply L_exten1; [trivial | move=> u uI; rewrite /f1 /g1 - card_multA //].
rewrite -/I p1.
move=> <-.
rewrite P3.
rewrite - cardinal_distrib_prod_sum2 card_multC.
rewrite (Bsucc_rw nB) power_of_sum2 power_x_1 /X; [apply refl_equal| fprops].
Qed.


Definition set_of_incr_functions r r' :=
  (Zo (set_of_functions (substrate r) (substrate r'))
    (fun z => increasing_fun z r r')).
Definition set_of_strict_incr_functions r r' :=
  (Zo (set_of_functions (substrate r) (substrate r'))
    (fun z => strict_increasing_fun z r r')).

Lemma cardinal_set_of_increasing_functions1: forall r r' f,
  total_order r -> strict_increasing_fun f r r' ->
  (order_morphism f r r' & cardinal (substrate r) = cardinal (image_of_fun f)).
Proof.
move=> r r' f tor si.
move: (total_order_increasing_morphism tor si) => mor; split =>//.
move: mor=> [_ [_ [injf [sr _]]]].
move: (restriction_to_image_bijective injf)=> bj.
rewrite sr; aw; exists (restriction_to_image f).
split =>//; rewrite /restriction_to_image /restriction2; aw; ee.
Qed.

Lemma cardinal_set_of_increasing_functions2: forall r r',
  total_order r -> total_order r' ->
  finite_set (substrate r) -> finite_set (substrate r') ->
  bijection (BL (fun z => range (graph z))
     (set_of_strict_incr_functions r r')
    (subsets_with_p_elements (cardinal (substrate r)) (substrate r'))).
Proof.
move=> r r' tor tor' fsr fsr'.
rewrite /set_of_strict_incr_functions.
set src:=Zo _ _.
set trg:= subsets_with_p_elements _ _.
set f:= BL(fun z=> range (graph z)) src trg.
have ta: transf_axioms (fun z=> range (graph z)) src trg.
  move=> z; rewrite /src Z_rw.
  aw; move=> [[fz [sc tc]] siz].
  move: (image_of_fun_range fz)=> imc.
  move:(cardinal_set_of_increasing_functions1 tor siz)=> [morz cs].
  rewrite /trg/subsets_with_p_elements Z_rw powerset_inc_rw cs -tc imc.
  by ee;apply f_range_graph.
apply bl_bijective =>//.
  move=> u v; rewrite /src Z_rw Z_rw; aw.
  move => [[fu [su tu]] siu][[fv [sv tv]] siv] sr.
  apply function_exten =>//; [rewrite sv // | rewrite tv //|].
  move=> x xsu; ex_middle Wx.
  move: (finite_set_torder_worder tor fsr) => [or wor].
  set Z := Zo (substrate r) (fun x=> W x u <> W x v).
  have p1: sub Z (substrate r) by apply Z_sub.
  have p2: nonempty Z by exists x; apply Z_inc => //; ue.
  move: (wor _ p1 p2)=> [y []]; aw; move => p3 p4.
  move: p3; rewrite /Z Z_rw -sv; move => [ysv Wy].
  have le: forall t, glt r t y -> W t u = W t v.
    move=> t ty; ex_middle Wt.
    have tZ: inc t Z by rewrite /Z; apply Z_inc=>//; order_tac.
    move:(iorder_gle1 (p4 _ tZ)) => yt.
    order_tac.
  move: (total_order_increasing_morphism tor siu) =>[_[_[[_ iu] _]]].
  move: (total_order_increasing_morphism tor siv) =>[_[_[[_ iv] _]]].
  rewrite su -sv in iu.
  have: inc (W y u) (range (graph v)).
     rewrite -sr; apply inc_W_range_graph=>//; rewrite su -sv //.
  rewrite range_inc_rw //; move=> [a [asv Wa]].
  have ya: glt r y a.
    move: (@total_order_pr2 r a y tor); rewrite -sv; move=> aux.
    case (aux asv ysv) => h.
      rewrite - (le _ h) in Wa.
      rewrite (iu _ _ ysv asv Wa) in h; red in h; intuition.
      split =>//; move => h'; rewrite -h' in Wa; contradiction.
  have: inc (W y v) (range (graph u)).
     rewrite sr; apply inc_W_range_graph=>//.
  rewrite range_inc_rw //; rewrite su -sv; move=> [b [bsv Wb]].
  have yb: glt r y b.
    move: (@total_order_pr2 r b y tor); rewrite -sv; move=> aux.
    case (aux bsv ysv) => h.
      rewrite (le _ h) in Wb.
      rewrite (iv _ _ ysv bsv Wb) in h; red in h; intuition.
      split =>//; move => h'; rewrite -h' in Wb; by elim Wy.
  move: siv =>[_ [_ [_ [_ [_ siv]]]]]; move: (siv _ _ ya); rewrite -Wa => h1.
  move: siu =>[_ [_ [w[_ [_ siu]]]]]; move: (siu _ _ yb); rewrite -Wb => h2.
  order_tac.
have or':order r' by move: tor' => [tor' _].
move=> y; rewrite /trg/subsets_with_p_elements Z_rw; aw; move => [ys ey].
move: (total_order_sub tor' ys).
set (r'' := induced_order r' y).
move=> tor''.
have sr'': (substrate r'' = y) by rewrite /r''; aw.
have e1: ( (substrate r) \Eq (substrate r'')) by rewrite sr''; eqsym.
move: (isomorphism_worder_finite tor tor'' fsr e1) => [[g gi] _].
move: (gi)=> [_ [_ [bg [sg [tg ig]]]]].
have cc: ((canonical_injection y (substrate r')) \coP g).
  red; ee ; [fct_tac | rewrite / canonical_injection -sr''; aw].
set (x:= (canonical_injection y (substrate r')) \co g).
have fg: (is_function x) by rewrite /x; fct_tac.
have sx: (source x = substrate r) by rewrite /x sg; aw.
have tx: target x = substrate r'. rewrite /x/canonical_injection; aw.
have gp: forall a, inc a (source g) -> inc (W a g) y.
   rewrite -sr'' tg; move=> a asg; apply inc_W_target =>//; fct_tac.
have xp: forall a, inc a (source g) -> W a x = W a g.
   move=> a ag; rewrite /x; aw; rewrite ci_W //; apply gp=>//.
have gxp: forall a, inc a (source g) -> inc (W a x) y.
   move=> a ax; rewrite xp //; apply gp=>//.
move: bg => [[fcg ing] sug].
have xsr: inc x src.
  rewrite /src; apply Z_inc; aw; ee.
  red; ee.
    by move: tor=> [tor _].
  move=> a b [ab nab].
  have asx: (inc a (source g)) by rewrite -sg; order_tac.
  have bsx: (inc b (source g)) by rewrite -sg; order_tac.
  rewrite (xp _ asx) (xp _ bsx).
  split =>//.
   move: ab;rewrite (ig _ _ asx bsx) /r'' => h.
  apply (iorder_gle1 h).
  dneg nW; apply ing =>//.
ex_tac.
set_extens t; rewrite range_inc_rw //.
  move=> [u []]; rewrite sx sg => us tw; rewrite tw; apply gxp =>//.
rewrite -sr'' tg=> tt; move: (surjective_pr2 sug tt)=> [a [asx av]].
rewrite sx sg -av -(xp _ asx).
ex_tac.
Qed.

Lemma cardinal_set_of_increasing_functions: forall r r',
  total_order r -> total_order r' ->
  finite_set (substrate r) -> finite_set (substrate r') ->
  cardinal (set_of_strict_incr_functions r r')
  = binom (cardinal (substrate r')) (cardinal (substrate r)).
Proof.
move=> r r' tor tor' fsr fsr'.
move: (cardinal_set_of_increasing_functions2 tor tor' fsr fsr').
set src := set_of_strict_incr_functions r r'.
set trg := subsets_with_p_elements _ _.
move=> h.
have -> :cardinal src = cardinal trg.
    aw;exists (BL (fun z : Set => range (graph z)) src trg); ee; aw.
rewrite /trg.
symmetry;set n := cardinal _; set p := cardinal _.
have nB: inc n Bnat by rewrite inc_Bnat /n; apply fsr'.
have pB: inc p Bnat by rewrite inc_Bnat /p; apply fsr.
have cE: cardinal (substrate r') = n by [].
by rewrite (subsets_with_p_elements_pr nB pB cE).
Qed.

Lemma increasing_prop0 : forall p f r, inc p Bnat -> order r ->
  (forall i, i <=c p -> inc (f i) (substrate r)) ->
  (forall n, n <c p -> gle r (f n) (f (succ n))) ->
  (forall i j, i <=c j -> j <=c p ->
    gle r (f i) (f j)).
Proof.
move=> p f r pB or als hyp i j ij jp.
have jB: inc j Bnat by Bnat_tac.
have iB: inc i Bnat by Bnat_tac.
move:(card_sub_pr0 jB iB ij).
set k := (card_sub j i); move=> [kB ikj]; rewrite -ikj.
have ckk: cardinal_le k k by fprops.
set (pr:= fun n => (n <=c k) -> gle r (f i) (f (card_plus i n))).
apply (cardinal_c_induction (r:= pr)) =>//.
  move=> aux; aw; [ order_tac; apply als; co_tac | fprops ].
move => n nB prn h.
have le2: n <c (succ n) by srw; fprops.
have le4: (card_plus i n) <c p.
  have le3: n <c k by co_tac.
  move: (finite_sum3_lt nB iB kB le3).
  rewrite card_plusC (card_plusC k i) ikj => le4; co_tac.
move: (hyp _ le4).
have -> : (succ (i +c n) = (i +c (succ n))) by rewrite -plus_via_succ //.
move=> r1;move: le2 => [le2 _].
have le3: n <=c k by co_tac.
move: (prn le3) => r2; order_tac.
Qed.

Lemma strict_increasing_prop0 : forall p f r, inc p Bnat -> order r ->
  (forall n, n <c p -> glt r (f n) (f (succ n))) ->
  (forall i j, i <c j -> j <=c p ->
    glt r (f i) (f j)).
Proof.
move=> p f r pB or hyp i j ij jp.
have jB: inc j Bnat by Bnat_tac.
have iB: inc i Bnat by move: ij => [ij _ ];Bnat_tac.
have ip: i <c p by co_tac.
move: ij; rewrite - lt_n_succ_le0 ; fprops; move => ij.
have siB: inc (succ i) Bnat by hprops.
move:(card_sub_pr0 jB siB ij).
set k := (j -c (succ i)); move=> [kB ikj]; rewrite -ikj.
have ckk: k <=c k by fprops.
set pr:= fun n => (n <=c k) -> glt r (f i) (f ((succ i) +c n)).
apply (cardinal_c_induction (r:= pr)) =>//.
  move=> aux; aw; [ apply hyp =>// | fprops].
move => n nB prn h.
set m := succ (i +c n).
have mv: m = ((succ i) +c n).
  rewrite card_plusC /m card_plusC plus_via_succ//.
have <-: (succ m = card_plus (succ i) (succ n)).
   rewrite mv (Bsucc_rw iB) (Bsucc_rw nB) card_plusA.
   rewrite Bsucc_rw; hprops.
have le2: n <c (succ n) by srw; fprops.
have mp: m <c p.
  have mj: m <c j. rewrite /m -ikj.
    have ltkn: n <c k by co_tac.
    move: (finite_sum3_lt nB siB kB ltkn).
  by rewrite -/m card_plusC -mv (card_plusC k (succ i)).
  co_tac.
move: (hyp _ mp) => r1.
have ltkn: n <c k by co_tac.
move: ltkn => [ltkn _];move: (prn ltkn).
rewrite -mv; move => [r2 _]; order_tac.
Qed.

Lemma increasing_prop : forall p f r, inc p Bnat -> is_function f ->
  source f = interval_co_0a (succ p) -> order r -> substrate r = target f ->
  (forall n, n <c p -> gle r (W n f) (W (succ n) f)) ->
  increasing_fun f (interval_Bnato 0c p) r.
Proof.
move=> p f r pB ff sf or sr sge.
have zB: inc 0c Bnat by fprops.
have bi:interval_cc_0a p = source f.
   by rewrite sf; apply interval_co_cc =>//.
move: (worder_total (interval_Bnato_worder zB pB))=> [sor _ ].
red; ee.
  rewrite interval_Bnato_substrate //.
move=> x y; rewrite interval_Bnato_related //; move => [xs [ys lxy]].
move: ys;rewrite interval_Bnat_pr0 =>// => yp.
have aux: (forall i, cardinal_le i p -> inc (W i f) (substrate r)).
  move=> i ip; rewrite sr; apply inc_W_target =>//.
  rewrite sf interval_co_0a_pr3 //.
apply (increasing_prop0 pB or aux sge lxy yp).
Qed.

Lemma strict_increasing_prop : forall p f r, inc p Bnat -> is_function f ->
  source f = interval_co_0a (succ p) -> order r -> substrate r = target f ->
  (forall n, n <c p -> glt r (W n f) (W (succ n) f)) ->
  (injection f &
    strict_increasing_fun f (interval_Bnato 0c p) r).
Proof.
move => p f r pB ff sf or sr prop.
have zB: inc 0c Bnat by fprops.
have bi:interval_cc_0a p = source f.
   by rewrite sf; apply interval_co_cc =>//.
move: (worder_total (interval_Bnato_worder zB pB)) => tor.
have sic: (strict_increasing_fun f (interval_Bnato 0c p) r).
  red; ee.
      by move: tor => [sor _].
    rewrite interval_Bnato_substrate //.
    move=> x y [cle xne]. move: cle; rewrite interval_Bnato_related //.
    move => [xs [ys lxy]].
    move: ys;rewrite interval_Bnat_pr0 =>// => yp.
    have ltxy: cardinal_lt x y by split.
    apply (strict_increasing_prop0 pB or prop ltxy yp).
split=>//.
move: (total_order_increasing_morphism tor sic).
rewrite /order_morphism; intuition.
Qed.

Lemma strict_increasing_prop1: forall f p,
  inc p Bnat -> (forall i, i <c p -> inc (f i) Bnat)
  -> (forall i j, i <c j -> j <c p -> (f i) <c (f j)) ->
  (forall i, i <c p -> i <=c (f i)).
Proof.
move=> f p pB ali hi i lip.
have iB: (inc i Bnat) by move: lip=> [leip _]; Bnat_tac.
move: i iB lip; apply: cardinal_c_induction.
  move=> aux;apply zero_smallest.
  move: (ali _ aux); fprops.
move=> n B pn aux.
have le2: n <c (succ n) by srw; fprops.
move: (hi _ _ le2 aux) => le3.
have le4: n <c p by co_tac.
move: (pn le4) => aux1.
srw;[ co_tac | fprops | apply ali=>//].
Qed.

Lemma increasing_prop1 : forall p f, inc p Bnat ->
  (forall i, i <=c p -> inc (f i) Bnat) ->
  (forall n, n <c p -> (f n) <=c (f (succ n))) ->
  (forall i j, i <=c j -> j <=c p -> (f i) <=c (f j)).
Proof.
move => p f pB p1 p2 i j ij jp.
move: (Bnat_order_worder) => [or _].
set r := Bnat_order.
have q1: (forall i, i <=c p -> inc (f i) (substrate r)).
  by rewrite Bnat_order_substrate.
have q2:(forall n, n <c p -> gle r (f n) (f (succ n))).
  move=> n np; rewrite Bnat_order_le; red; ee.
  by apply p1; move: np => [np _].
  apply p1; rewrite lt_n_succ_le0 //; fprops.
  by move: np=> [[cn _] _].
move: (increasing_prop0 pB or q1 q2 ij jp).
rewrite Bnat_order_le /Bnat_le; intuition.
Qed.

Lemma strict_increasing_prop2: forall f p,
  inc p Bnat -> (forall i, i <c p -> inc (f i) Bnat)
  -> (forall i j, i <c j -> j <c p -> (f i) <c (f j)) ->
  (forall i j, i <=c j -> j <c p -> ((f i) -c i) <=c ((f j) -c j)).
Proof.
move=> f p pB ali hyp.
case (equal_or_not p 0c).
  move => -> i j ij jp; elim (zero_smallest1 jp).
move=> pnz.
move:(predc_pr pB pnz) => [qB ps].
set q := predc p.
have fq: finite_c (predc p) by rewrite -inc_Bnat.
set (g:= fun i=> (card_sub (f i) i)).
have q1: (forall i : Set, cardinal_le i q -> inc (g i) Bnat).
  move=> i ip;
  have iB: inc i Bnat by Bnat_tac.
  rewrite /g; apply BS_sub =>//.
  apply ali; rewrite ps // lt_is_le_succ //.
have q2:forall n : Set, n <c q -> (g n) <=c (g (succ n)).
  move=> n np.
  have nB: inc n Bnat by move: np => [np _]; Bnat_tac.
  have nsB: inc (succ n) Bnat by hprops.
  rewrite /g.
  move:(strict_increasing_prop1 pB ali hyp) => hy.
  have p1: (succ n) <c p.
    by rewrite ps lt_is_le_succ // lt_n_succ_le0 //; fprops.
  have p2: n <c p.
     move: (is_le_succ nB)=> aux.
     co_tac.
  move: (hy _ p1) (hy _ p2)=> qp1 qp2.
  move: (ali _ p1)(ali _ p2) => s1 s2.
  move:(card_sub_pr0 s1 nsB qp1)(card_sub_pr0 s2 nB qp2).
  set a:= (f (succ n)) -c (succ n).
  set b := ((f n) -c n).
  move=> [aB r1][bB r2].
  have nsn: n <c (succ n) by apply is_lt_succ; fprops.
  move: (hyp _ _ nsn p1);rewrite -r1 -r2.
  have->: ((succ n) +c a = succ (n +c a)).
    by rewrite card_plusC (plus_via_succ _ nB) card_plusC.
  rewrite lt_is_le_succ //.
    apply plus_le_simplifiable =>//.
  hprops.
move=> i j ij jp.
apply (increasing_prop1 qB q1 q2 ij).
by move: jp; rewrite {1} ps lt_is_le_succ.
Qed.

Lemma strict_increasing_prop3: forall f p n,
  inc p Bnat -> inc n Bnat -> (forall i, i <c p -> inc (f i) Bnat)
  -> (forall i j, i <c j -> j <c p -> (f i) <c (f j)) ->
  (forall i, i <c p -> (f i) <c (n +c p)) ->
  (forall i, i <c p -> ((f i) -c i) <=c n).
Proof.
move=> f p n pB nB ali fi fb.
move=> i ip.
case (equal_or_not p 0c) => h.
  rewrite h in ip;elim (zero_smallest1 ip).
move: (predc_pr pB h) => [qB psq].
move:ip; rewrite psq ; srw; move => iq.
have qp: ((predc p) <c p) by rewrite {2} psq; srw; fprops.
apply cardinal_le_transitive with (card_sub (f (predc p)) (predc p)).
  apply (strict_increasing_prop2 pB ali fi iq qp).
move: (strict_increasing_prop1 pB ali fi qp) => aa.
have fB: (inc (f (predc p)) Bnat) by apply ali =>//.
move: (card_sub_pr aa) => csp.
move: (fb _ qp); rewrite - csp.
set (r:= ((f (predc p)) -c (predc p))).
have rB: inc r Bnat by rewrite /r;hprops.
have ->: (n +c p = succ ( (predc p) +c n)).
  by rewrite {1} psq (plus_via_succ _ qB) card_plusC.
move=> hh.
have sB: inc ((predc p) +c n) Bnat by hprops.
apply plus_le_simplifiable with (predc p); hprops.
rewrite card_sub_pr - ? (lt_is_le_succ _ sB) //;apply sum_increasing3; fprops.
Qed.

Lemma cardinal_set_of_increasing_functions3: forall n p,
  inc n Bnat -> inc p Bnat ->
  cardinal (set_of_incr_functions (interval_Bnatco p)
         (interval_Bnato 0c n))
   = binom (n +c p) p.
Proof.
move=> n p nB pB.
set r := (interval_Bnatco p).
set r' := (interval_Bnato 0c n).
set (E1:=interval_co_0a p).
have s1: E1= substrate r by rewrite /E1 interval_Bnatco_substrate.
set E2:=interval_cc_0a n.
have s2: (E2= substrate r') by rewrite /E2 interval_Bnato_substrate; fprops.
set np := n +c p.
have npB: inc np Bnat by rewrite /np; hprops.
set (r'':= interval_Bnatco np).
set (E3:=interval_co_0a np).
have s3: E3= substrate r'' by rewrite /E3 interval_Bnatco_substrate.
set Q1:= set_of_incr_functions r r'.
set Q2:= set_of_strict_incr_functions r r''.
set (P1:= fun x => is_function x & source x = E1 & target x = E2 &
    increasing_fun x r r').
set (P2:= fun x => is_function x & source x = E1 & target x = E3 &
    strict_increasing_fun x r r'').
have HQ1: forall x, inc x Q1 = P1 x.
  move=> x.
  rewrite /Q1 /P1 s1 s2 /set_of_incr_functions Z_rw;aw.
  apply iff_eq; intuition.
have HQ2: forall x, inc x Q2 = P2 x.
  move=> x.
  rewrite /Q2 /P2 s1 s3 /set_of_incr_functions Z_rw;aw.
  apply iff_eq; intuition.
have E1B :sub E1 Bnat by rewrite /E1; apply sub_interval_co_0a_Bnat.
have E2B :sub E2 Bnat by rewrite /E2; apply sub_interval_Bnat.
have E3B :sub E3 Bnat by rewrite /E3; apply sub_interval_co_0a_Bnat.
set (subi:= fun f => BL (fun i=> card_sub (W i f) i) E1 E2).
have Hi0:forall f, inc f Q2 -> (
    (forall i, i <c p -> inc (W i f) Bnat) &
    (forall i j, i <c j -> j <c p -> (W i f) <c (W j f)) &
    (forall i, inc i E1 -> (i <=c (W i f)))).
  move=> f; rewrite HQ2; move=> [ff [sf [tf sif]]].
  have aux: (forall i, i <c p -> inc (W i f) Bnat).
    move=> i ilp; apply E3B =>//; rewrite -tf; apply inc_W_target =>//.
     rewrite sf /E1; srw.
  have aux2:(forall i j, i <c j -> j <c p -> (W i f) <c (W j f)).
    move=> i j ij jp.
    have isf: (inc i (source f)) by rewrite sf /E1; srw; co_tac.
    have jsf: (inc j (source f)) by rewrite sf /E1; srw; co_tac.
    have lij: glt r i j.
       move: ij => [ij nij]; rewrite /r; split=>//.
       rewrite interval_Bnatco_related; ee.
    move: sif => [_[_[_[_[_ h]]]]]; move: (h _ _ lij) => [lW nW].
    split =>//; move: lW; rewrite /r'' interval_Bnatco_related//; intuition.
  ee.
  move=> i; rewrite /E1; srw. apply: (strict_increasing_prop1 pB aux aux2).
have Hi3:forall f, inc f Q2 ->
    transf_axioms (fun i => (W i f) -c i) E1 E2.
  move=> f fQ2; move: (Hi0 _ fQ2) => [p1 [p2 p3]].
  move: (strict_increasing_prop2 pB p1 p2) => p4.
  have p5: (forall i, i <c p-> (W i f) <c (n +c p)).
    move=> i ip.
    move: fQ2; rewrite HQ2 /P2; move => [ff [sf [tf vf]]].
    have : (inc (W i f) E3).
      by rewrite -tf; apply inc_W_target =>//; rewrite sf /E1; srw.
    rewrite /E3; srw.
  move: (strict_increasing_prop3 pB nB p1 p2 p5) => p6.
  move=> i; rewrite /E1 /E2; rewrite interval_Bnat_pr0; srw; apply p6.
have Hi1:forall f, inc f Q2 -> inc (subi f) (set_of_functions E1 E2).
  by move=> f fQ2; aw; rewrite /subi; ee; aw; apply bl_function; apply Hi3.
have Hi2:forall f, inc f Q2 -> inc (subi f) Q1.
  move=> f fQ2; rewrite /Q1; apply Z_inc; first by rewrite -s1 -s2; apply Hi1.
  move: (Hi3 _ fQ2) => ta.
  rewrite /subi;red; aw;ee.
        apply bl_function =>//.
      by move: (interval_Bnatco_worder pB) => [or _].
    by move: (interval_Bnato_worder inc0_Bnat nB) => [or _].
  move=> x y xy.
  have xsf: (inc x E1) by rewrite s1; order_tac.
  have ysf: (inc y E1) by rewrite s1; order_tac.
  move: (Hi0 _ fQ2) => [p1 [p2 p3]].
  move:(strict_increasing_prop2 pB p1 p2) => p4.
  aw.
  have yp: y <c p by move: ysf; rewrite /E1; srw.
  have xy': x <=c y.
     rewrite /r interval_Bnatco_related in xy; intuition.
  move: (p4 _ _ xy' yp); rewrite /r' interval_Bnato_related; ee.
set (G:= BL (fun f=> subi f) Q2 Q1).
have Hi4:is_function G by rewrite / G; apply bl_function.
set (addi:= fun f => BL (fun i=> (W i f) +c i) E1 E3).
have pa:forall f, inc f (set_of_functions E1 E2) ->
    transf_axioms (fun i => (W i f) +c i) E1 E3.
  move=> f; aw; move=> [ff [sf tf]] t tE1.
  have: (inc (W t f) E2) by rewrite -tf;apply inc_W_target=>//; ue.
  move: tE1; rewrite /E1/E2/E3 interval_Bnat_pr0; srw.
  move => p1 p2; rewrite /np; apply finite_sum2_lt =>//.
    Bnat_tac.
  move: p1 => [p1 _]; Bnat_tac.
have pa2: (forall f, inc f (set_of_functions E1 E2) ->
    inc (addi f) (set_of_functions E1 E3)).
  move=> f fE; rewrite /addi; aw; ee.
  by apply bl_function; apply pa.
have pa3: forall f, inc f Q1 -> inc (addi f) Q2.
  move=> f; rewrite /Q1 Z_rw; rewrite -s1 -s2; move=>[fe1e2 incf].
  rewrite HQ2 /P2; move: (pa2 _ fe1e2); aw; move => [fa [sa ta]]; ee.
  red; ee; try ue.
      by move: (interval_Bnatco_worder pB) => [or _].
    by move: (interval_Bnatco_worder npB) => [or _].
  move=> x y [xy nxy].
  have xe1: inc x E1 by rewrite s1; order_tac.
  have ye1: inc y E1 by rewrite s1; order_tac.
  move: (pa _ fe1e2) => ta1.
  rewrite /addi bl_W // bl_W //.
  move: incf => [_[_[_[_ [_ h]]]]]; move: (h _ _ xy); rewrite /r'/r'' /glt.
  have zB: inc 0c Bnat by fprops.
  rewrite interval_Bnatco_related // interval_Bnato_related2 //.
  move=> [p1 [p2 p3]].
  have ltxy: cardinal_lt x y.
    split =>//; move: xy; rewrite /r interval_Bnatco_related //; intuition.
  have aux: cardinal_le (W x f) n by co_tac.
  have: ( (W x f) +c x) <c ( (W y f) +c y).
     apply finite_sum2_lt =>//; try Bnat_tac; try apply E1B =>//.
  move=> [p4 p5];ee.
  rewrite /np; apply finite_sum2_lt =>//; [Bnat_tac | apply E1B=>//|].
  move: ye1; rewrite /E1; srw.
set (F:= BL (fun f=> (addi f)) Q1 Q2).
have ff: is_function F by rewrite /F;apply bl_function.
have cGF: composable G F by rewrite /G/F; red; ee; aw.
have cFG: composable F G by rewrite /G/F; red; ee; aw.
have c1i: (compose G F = identity (source F)).
  apply function_exten; aw; [fct_tac |apply identity_function | | ].
     rewrite /G/F; aw.
  move => x xsf;rewrite identity_W //.
  move: xsf ;rewrite /F; rewrite {1} corresp_source => xQ1.
  have aQ2: inc (addi x) Q2 by apply pa3.
  aw; rewrite /G; aw.
  rewrite /addi/subi.
  move: (xQ1); rewrite HQ1; move=> [fx [sx [tx icx]]].
  apply function_exten; aw =>//.
     apply bl_function; apply Hi3; apply pa3 =>//.
  move: (Hi3 _ aQ2)=> ta2.
  move=> a aE1; aw.
    apply card_sub_pr1; last by apply E1B.
    apply E2B;rewrite -tx//;apply inc_W_target=> //; ue.
  apply pa; aw; ee.
have c2i: (compose F G = identity (source G)).
  apply function_exten.
  - fct_tac.
  - apply identity_function.
  - rewrite /F/G; aw.
  - rewrite /F/G;aw.
  - rewrite /G; aw.
  - move=> x xQ2;rewrite identity_W //.
   have qQ1: inc (subi x) Q1 by apply Hi2.
   move: (xQ2); rewrite HQ2; move=> [fx [sx [tx icx]]].
   aw;rewrite /F; aw; rewrite /addi/subi.
   apply function_exten =>//.
   - apply bl_function; apply pa; apply Hi1=> //.
   - symmetry;aw.
   - symmetry; aw.
   - move: (Hi3 _ xQ2) => ta3.
     move:(pa _ (Hi1 _ xQ2)) => ta2.
     aw; move => a aE1; aw; apply card_sub_rpr.
     move: (Hi0 _ xQ2); intuition.
move: (bijective_from_compose cGF cFG c1i c2i) =>[bF [bG GF]].
have : (equipotent Q1 Q2) by exists F; rewrite /F;ee; aw .
rewrite - cardinal_equipotent; move=> ->.
have r1: (cardinal (substrate (interval_Bnatco np))) = np.
  by rewrite interval_Bnatco_substrate // cardinal_interval_co_0a1.
have r2: (cardinal (substrate (interval_Bnatco p))) = p.
  by rewrite interval_Bnatco_substrate // cardinal_interval_co_0a1.
rewrite /Q2 /r/r'' cardinal_set_of_increasing_functions /finite_set.
- rewrite r1 r2 //.
- apply worder_total; apply interval_Bnatco_worder =>//.
- apply worder_total; apply interval_Bnatco_worder =>//.
- rewrite r2 -inc_Bnat //.
- rewrite r1 -inc_Bnat //.
Qed.

Lemma increasing_compose: forall f g r r' r'',
  increasing_fun f r r' -> increasing_fun g r' r'' ->
  (g \coP f &
    (forall x, inc x (source f) -> W x (g \co f) = W (W x f) g) &
    increasing_fun (g \co f) r r'').
Proof.
move=> f g r r' r'' [ff [or [or' [sf [tf icf]]]]][fg [_ [or'' [sg [tg icg]]]]].
have cgf: ( g \coP f) by red; ee; ue.
have p:(forall x, inc x (source f) -> W x (g \co f) = W (W x f) g).
  move=> x xsf; aw.
ee; red; ee; try fct_tac; aw.
move => x y xy.
   have xsf: inc x (source f) by rewrite -sf; order_tac.
   have ysf: inc y (source f) by rewrite -sf; order_tac.
by rewrite p // p //; apply icg; apply icf.
Qed.

Lemma increasing_compose3: forall f g h r r' r'' r''',
  strict_increasing_fun f r r' -> increasing_fun g r' r'' ->
  strict_increasing_fun h r'' r''' ->
  let res:= (h \co g) \co f in
    (inc res (set_of_functions (source f) (target h)) &
      (forall x, inc x (source f) -> W x res = W (W (W x f) g) h) &
      increasing_fun res r r''').
Proof.
move=> f g h r r' r'' r''' sif ig sih.
move: (increasing_fun_from_strict sif) => icf.
move: (increasing_fun_from_strict sih) => ich.
move: (increasing_compose ig ich)=> [chg [ Whg igh]].
move: (increasing_compose icf igh)=> [chgf [ Whgf ighf]].
move=> res; rewrite /res;ee.
  aw; ee;fct_tac.
move=> x xsf; aw.
move: (increasing_compose icf ig)=> [[_ [ff sgtf]] _].
rewrite sgtf; apply inc_W_target =>//.
Qed.

Lemma cardinal_set_of_increasing_functions4: forall r r',
  let n := (cardinal (substrate r')) in
    let p := (cardinal (substrate r)) in
      total_order r -> total_order r' ->
      finite_set (substrate r) -> finite_set (substrate r') ->
      cardinal (set_of_incr_functions r r')
      = binom ((n +c p) -c 1c) p.
Proof.
move=> r r' n p tor tor' fsr fsr'.
have nB: inc n Bnat by move: fsr'; rewrite inc_Bnat /finite_set.
have pB: inc p Bnat by move: fsr; rewrite inc_Bnat /finite_set.
have or: order r by move: tor => [tor _ ].
have or': order r' by move: tor' => [tor' _ ].
case (equal_or_not n 0c) => nz.
  rewrite nz; aw; fprops.
  rewrite /set_of_incr_functions.
  have qe: substrate r'= emptyset by apply cardinal_nonemptyset; rewrite-nz//.
  rewrite qe.
  case (equal_or_not p 0c) => pz.
    rewrite pz card_sub_wrong; fprops; last first.
      move => h; move: (zero_smallest2 h); apply card_one_not_zero.
    rewrite binom00.
    have qe': substrate r = emptyset.
      by apply cardinal_nonemptyset; rewrite -pz//.
    rewrite qe'; set qq:=Zo _ _.
    move:empty_function_function => ef.
    have ->: qq = singleton empty_function.
      set_extens x;rewrite /qq Z_rw; aw.
        move=> [[fx [sx tx]] _].
        apply function_exten =>//; rewrite /empty_function ; aw.
        rewrite sx; move=> a ae; elim (emptyset_pr ae).
      move=> ->; rewrite /empty_function; aw;ee; red; aw;ee.
      move=> a b ab.
      have: (inc a (substrate r)) by order_tac.
      rewrite qe'; move=> ae; elim (emptyset_pr ae).
    by rewrite cardinal_singleton.
  move: (predc_pr pB pz) => [ppB sp].
  have ->: (card_sub p card_one = predc p) by rewrite prec_pr1 //.
  have ltp:cardinal_lt (predc p) p by rewrite {2} sp; apply is_lt_succ; fprops.
  rewrite binom_bad //.
  set qq:=Zo _ _ ; case (emptyset_dichot qq).
    by move=> ->; apply cardinal_emptyset.
  move=> [t tq]; move: tq; rewrite /qq Z_rw; aw.
  move=> [[ft [st tt]] _].
  case (emptyset_dichot (substrate r)) => h.
    by elim pz; rewrite /p h cardinal_emptyset.
  move: h=> [x xsr]; rewrite -st in xsr; move: (inc_W_target ft xsr).
  rewrite tt=>h; elim (emptyset_pr h).
move: (predc_pr nB nz) => [pnB sn].
set q := (predc n) in sn pnB.
have ->: ( (n +c p) -c 1c = q +c p).
  rewrite sn; apply (card_sub_pr2 (BS_plus pnB pB) inc1_Bnat).
  rewrite (plus_via_succ1 _ pnB) Bsucc_rw; hprops.
move: (finite_ordered_interval1 tor fsr) => [[f isf]_].
move: (finite_ordered_interval1 tor' fsr') => [[g isg] _].
move: (inverse_order_isomorphism isf)=> isif.
move: (inverse_order_isomorphism isg)=> isig.
rewrite- (cardinal_set_of_increasing_functions3 pnB pB).
rewrite -/n in isg. rewrite -/p in isf isif.
have ->: interval_Bnato 0c q = interval_Bnatco n.
  by rewrite /interval_Bnatco/interval_Bnato sn -interval_co_cc /interval_Bnat.
set Q1:= set_of_incr_functions _ _.
set Q2:= set_of_incr_functions _ _.
set f' := inverse_fun f.
set g' := inverse_fun g.
set ff:= BL (fun z => ((g \co z)) \co f') Q1 Q2.
have fg:is_function g by move: isg => [_[_[bg _]]]; fct_tac.
have fif:is_function f' by move: isif => [_[_[bg _]]]; fct_tac.
move: (order_isomorphism_increasing isg) => sig.
move: (order_isomorphism_increasing isif) => siif.
have tg: (substrate (interval_Bnatco n) = target g).
  by move: sig => [_ [_ [_ [_ [tg _]]]]].
have soif : (substrate (interval_Bnatco p) = source f')
  by move: siif => [_ [_ [_ [sf _]]]].
have tf': (substrate r = target f') by move: siif => [_ [_ [_ [_ [xx _]]]]].
have sg : (substrate r' = source g) by move: sig => [_ [_ [_ [xx _]]]].

have ta1:transf_axioms (fun z => (g \co z) \co f') Q1 Q2.
  move=> z;rewrite Z_rw; move => [zs iz].
  move: (increasing_compose3 siif iz sig); simpl; move=> [p1[p2 p3]].
  by rewrite /Q2 Z_rw; ee; rewrite tg soif.
have ffc3: (forall w, inc w Q1 -> W w (compose3function f' g)
    = (g \co w) \co f').
  move => w wQ; move: (wQ); rewrite /Q1 Z_rw; aw;move=> [[fw [sw tw]] siw].
  rewrite c3f_W //; ue.
have s_f: (surjection f') by move : isif=> [_ [_ [[_ ok] _]]].
have i_g: (injection g) by move : isg=> [_ [_ [[ok _] _]]].
have s3: sub Q1 (source (compose3function f' g)).
  rewrite /compose3function /Q1 /set_of_incr_functions -tf' -sg.
  aw; apply Z_sub.
aw; exists ff; rewrite /ff; aw; ee; apply bl_bijective => //.
  move=> u v u1 v1;rewrite -(ffc3 _ u1) -(ffc3 _ v1).
  by move: (c3f_injective s_f i_g) => [_ h]; apply h;apply s3.
move=> y yQ2.
have fg':is_function g' by move: isig => [_[_[bg _]]]; fct_tac.
have fif':is_function f by move: isf => [_[_[bg _]]]; fct_tac.
move: (order_isomorphism_increasing isig) => sig'.
move: (order_isomorphism_increasing isf) => siif'.
have tg': (substrate r'= target g') by move: isig => [_ [_ [_ [_ [xx _]]]]].
have soif' : (substrate r= source f) by move: siif' => [_ [_ [_ [xx _]]]].
have tf: (substrate (interval_Bnatco p)= target f)
  by move: siif' => [_ [_ [_ [_ [xx _]]]]].
have soig' :(substrate (interval_Bnatco n) = source g')
  by move: isig => [_ [_ [_ [xx _]]]].
set x := compose (compose g' y) f.
have xQ1: inc x Q1.
  move: yQ2;rewrite Z_rw; move => [zs iz].
  move: (increasing_compose3 siif' iz sig'); simpl; move=> [p1[p2 p3]].
  by rewrite /Q1 Z_rw; ee; rewrite tg' soif'.
ex_tac.
move: yQ2;rewrite Z_rw; move => []; aw; move=> [fy [sy ty]] _.
have b_f: (bijection f) by move : isf=> [_ [_ [ok _]]].
have b_g: (bijection g) by move : isg=> [_ [_ [ok _]]].
have c1: g' \coP y by red; ee; ue.
have sysf: source y = target f by ue.
have c2: y \coP f by red; ee; ue.
have fyg: is_function (y \co f) by fct_tac.
have tyg: target g = target (y \co f) by aw; ue.
have c4: g' \coP (y \co f) by red; ee; aw; ue.
rewrite /x (compose_assoc c1 c2).
rewrite -(compose_assoc (composable_f_inv b_g) c4).
rewrite (bij_right_inverse b_g) tyg (compose_id_left fyg).
rewrite /f' (compose_assoc c2 (composable_f_inv b_f)) (bij_right_inverse b_f).
by rewrite -sysf (compose_id_right fy).
Qed.

Lemma cardinal_lt20: 0c <c 2c.
Proof. rewrite -succ_one lt_is_le_succ; fprops. Qed.

Lemma binom_2plus: forall n, inc n Bnat ->
  binom (succ n) 2c = (n *c (succ n)) %/c 2c.
Proof. move=> n nB;rewrite binom2//.
Qed.

Lemma binom_2plus0: forall n, inc n Bnat ->
  binom (succ n) 2c = (binom n 2c) +c n.
Proof. move=> n nB.
have oB: inc card_one Bnat by fprops.
by rewrite -succ_one (binomSnSm nB oB) (binom1 nB).
Qed.

Lemma cardinal_pairs_lt: forall n, inc n Bnat ->
  cardinal (Zo (product Bnat Bnat)
    (fun z => 1c <=c (P z) & (P z) <c (Q z) & (Q z) <=c n)) =
  (binom n 2c).
Proof.
move=> n nB.
set (E:= interval_cc_1a n).
set T:=Zo _ _.
case (p_or_not_p (cardinal_le card_two n)); last first.
  case (cardinal_le_total_order2 cardinal2 (Bnat_is_cardinal nB)).
    by move=> h h'; contradiction.
  move=> h _; rewrite binom_bad //; try apply inc2_Bnat.
  case (emptyset_dichot T); first by move=> ->; rewrite cardinal_emptyset.
  move=> [p pT]; move: pT; rewrite /T Z_rw; move=> [a1 [a2 [a3 a4]]].
  have : cardinal_lt (Q p) card_two by co_tac.
  move: inc1_Bnat => oB.
  rewrite -succ_one lt_is_le_succ// => a5.
  have a6: cardinal_lt (P p) card_one by co_tac.
  co_tac.
move=> le2n.
have cE: cardinal E = n by rewrite cardinal_interval1a.
rewrite (subsets_with_p_elements_pr nB inc2_Bnat cE).
aw.
exists (BL (fun z=> (doubleton (P z) (Q z))) T
    (subsets_with_p_elements card_two E)).
aw; ee; apply bl_bijective.
- move=> z ; rewrite /T Z_rw; move=>[zp [le1p [[lepq npq] leqn]]].
   have pE: inc (P z) E.
    rewrite /E interval_Bnat_pr1b //; ee; co_tac.
   have qE: inc (Q z) E.
    rewrite /E interval_Bnat_pr1b //; ee; co_tac.
  rewrite /subsets_with_p_elements; apply Z_inc.
     by aw; move=> t; rewrite doubleton_rw; case; move=> ->.
  rewrite cardinal_doubleton //.
- move=> u v; rewrite /T Z_rw Z_rw; aw.
  move=> [[pu _] [_ [lt1 _]]] [[pv _] [_ [lt2 _]]] h.
  case (doubleton_inj h); move=> [e1 e2].
    apply pair_exten =>//.
    move: lt1;rewrite e1 e2; move => [lt1 _]; co_tac.
- move => y; rewrite/subsets_with_p_elements Z_rw; aw; move=> [yE cy2].
  have yB: sub y Bnat.
    apply sub_trans with E => //; rewrite /E; apply sub_interval_Bnat.
  have [a [b [p1 p2]]]:
    exists a, exists b, cardinal_lt a b & doubleton a b = y.
    move: (set_of_card_two cy2)=> [a [b [ab yab]]].
    have aE: inc a Bnat by apply yB; rewrite yab; fprops.
    have bE: inc b Bnat by apply yB; rewrite yab; fprops.
    case(cardinal_le_total_order2 (Bnat_is_cardinal aE)(Bnat_is_cardinal bE)).
      move=> h; exists a; exists b; ee; split=>//.
    by move=> h; exists b; exists a; ee; rewrite doubleton_symm.
  exists (J a b); aw; ee.
  rewrite /T; apply Z_inc.
  have aE: inc a Bnat by apply yB; rewrite -p2; fprops.
  have bE: inc b Bnat by apply yB; rewrite -p2; fprops.
  aw; fprops.
  aw.
  have: inc a E by apply yE; rewrite -p2; fprops.
  have: inc b E by apply yE; rewrite -p2; fprops.
  rewrite /E (interval_Bnat_pr1b _ nB) (interval_Bnat_pr1b _ nB).
  intuition.
Qed.

Lemma cardinal_pairs_le: forall n, inc n Bnat ->
  cardinal(Zo (product Bnat Bnat)
    (fun z=> 1c <=c (P z) & (P z) <=c (Q z) & (Q z) <=c n)) =
     (binom (succ n) 2c).
Proof.
move=> n nB.
rewrite (binom_2plus0 nB).
rewrite - cardinal_pairs_lt //.
set E1 := Zo _ _; set E2 := Zo _ _.
have s21: sub E2 E1.
  move=> t;rewrite /E1 /E2 Z_rw Z_rw /cardinal_lt; intuition.
rewrite -(cardinal_complement s21).
suff:(cardinal (complement E1 E2) = n) by move => ->.
set (T:= interval_Bnat card_one n).
have ->: n = cardinal T by rewrite cardinal_interval1a.
aw; exists (BL P (complement E1 E2) T); aw; ee.
have cp: forall x, inc x (complement E1 E2) -> P x = Q x.
  move=> x; srw;rewrite /E1/E2 Z_rw Z_rw ;move=> [[xp [le1 [le2 le]]] h].
  ex_middle npq; elim h; rewrite /cardinal_lt;intuition.
apply bl_bijective.
- move=> x xc; move: (cp _ xc); move: xc; srw.
   rewrite /E1 /T interval_Bnat_pr1b // Z_rw; move=> [[xp [le1 [le2 le]]] ns].
   move=> h; rewrite -h in le; ee.
- move => u v uc vc; move: (cp _ vc) (cp _ uc) => h1 h2.
   move: uc vc; srw. rewrite /E1 Z_rw Z_rw. aw.
   move=> [[[pu _] _] _] [[[pv _] _] _] Puv.
   by apply pair_exten => //; rewrite - h1 -h2 Puv.
- move=> y; rewrite /T interval_Bnat_pr1b //; move => [l1y lyn].
  have yB: inc y Bnat by Bnat_tac.
  exists (J y y); aw;ee; srw;rewrite /E1 /E2 Z_rw Z_rw; aw;ee.
  by move => [_ [_ [[ _ bad] _]]]; elim bad.
Qed.

Lemma sum_of_i: forall n, inc n Bnat ->
  cardinal_sum (L (interval_co_0a n) id) =
   binom n 2c.
Proof.
apply cardinal_c_induction.
  rewrite binom_bad; fprops.
  rewrite trivial_cardinal_sum =>//; bw; apply emptyset_interval_00.
  apply cardinal_lt20.
move=> n nB hr.
have aux: forall a : Set, inc a Bnat -> is_cardinal a by move=> a aB; fprops.
rewrite -induction_on_sum // hr binom_2plus0 //.
Qed.

Lemma fct_sum_const1: forall f n m,
  inc n Bnat -> (forall i, i <c n -> f i = m) ->
  cardinal_sum (L (interval_co_0a n) f) = n *c m.
Proof.
move=> f n m nB p.
set b := interval_co_0a n.
have <-: (cst_graph b m = L b f).
   rewrite /cst_graph; apply L_exten1 => //.
   by move=> x; rewrite /b; srw; symmetry; apply p.
rewrite sum_of_same1 card_multC.
apply card_mult_pr2=>//.
rewrite /b; rewrite cardinal_interval_co_0a1 //.
symmetry;apply cardinal_of_cardinal; fprops.
Qed.

Lemma sum_of_i2: forall n, inc n Bnat ->
  cardinal_sum (L (interval_Bnat 1c n) id) =
  (binom (succ n) 2c).
Proof.
move => n nB.
rewrite - cardinal_pairs_le //.
set (E:= interval_Bnat card_one n).
move: inc1_Bnat => oB.
set X:=Zo _ _ .
set(f:= L E (fun k => Zo X (fun z => Q z = k))).
have p1: X = unionb f.
  rewrite /f; set_extens x;srw ;bw.
    move => xX; move: (xX); rewrite Z_rw; move=> [xs [le1 [le2 le3]]].
    have qE:inc (Q x) E by rewrite /E interval_Bnat_pr1b //; ee; co_tac.
    exists (Q x); bw; split =>//; apply Z_inc=>//.
  move=> [y [yE]]; bw; rewrite Z_rw;intuition.
have p2: f= disjoint_union_fam (L E (fun i => interval_Bnat card_one i)).
  rewrite /f /disjoint_union_fam; bw; apply L_exten1 => //.
  move=> x xE; move: (xE); rewrite /E interval_Bnat_pr1b//; move=> [ox xn].
  have xB: inc x Bnat by Bnat_tac.
  bw; set_extens t; rewrite Z_rw /X Z_rw; aw.
    move => [[[pt [_ qB]] [oP [pq _]]] qt]; rewrite -qt interval_Bnat_pr1b; ee.
  rewrite interval_Bnat_pr1b //; move=> [pt [[le1 le2] le3]]; rewrite le3.
  intuition; Bnat_tac.
have ->: X = disjoint_union (L E (fun i => interval_Bnat card_one i)).
  by rewrite p1 p2 /disjoint_union.
rewrite /cardinal_sum;aw;apply equipotent_disjoint_union1; gprops; bw.
move => i iE; bw.
move: iE; rewrite /E interval_Bnat_pr1b //; move => [_ lin].
have iB: inc i Bnat by Bnat_tac.
rewrite - cardinal_equipotent cardinal_interval1a //.
apply cardinal_of_cardinal; fprops.
Qed.

Lemma sum_of_i3: forall n, inc n Bnat ->
  cardinal_sum (L (interval_co_0a n) id) =
   binom n 2c.
Proof.
move=> n nB.
move: cardinal_lt20 => lt20.
case (equal_or_not n 0c) => nz.
  rewrite nz trivial_cardinal_sum ? binom_bad; fprops.
  bw; apply emptyset_interval_00.
move: (predc_pr nB nz) => [pB nsp].
rewrite nsp binom2 //.
set p := predc n in pB nsp |- *.
set sn:= cardinal_sum _ .
set aux:= (p *c (succ p)).
suff: (sn +c sn = aux).
  move => h.
  have aB: inc aux Bnat by rewrite /aux; hprops.
  have iB: inc sn Bnat.
    have cn: is_cardinal sn by rewrite /sn/cardinal_sum; fprops.
    rewrite -h in aB; apply (Bnat_in_sum cn aB).
  rewrite - two_times_n in h.
  apply (BNdivides_pr2 aB inc2_Bnat iB card_two_not_zero (sym_equal h)).
move: (sum_of_sums (fun i => i) (fun i=> (p -c i))
     (interval_co_0a (succ p))).
rewrite nsp in nB.
have fim: (forall i : Set, i <c (succ p) ->
  (fun i : Set => i +c (p -ci)) i = p).
  move=> i ilp;rewrite lt_is_le_succ // in ilp.
  apply card_sub_pr =>//; Bnat_tac.
  rewrite (fct_sum_const1 nB fim).
rewrite card_multC /aux -/sn; move => <-; apply f_equal.
rewrite /sn.
rewrite fct_sum_rev //.
Qed.

Lemma sof_sum_eq_equi: forall F n, inc n Bnat ->
   (set_of_functions_sum_eq F n) \Eq (set_of_graph_sum_eq F n).
Proof.
move=> F n nB.
exists (BL graph (set_of_functions_sum_eq F n) (set_of_graph_sum_eq F n)).
aw; ee; apply bl_bijective.
    move => t ts; rewrite /set_of_graph_sum_eq; aw; ex_tac.
  move=> u v; rewrite /set_of_functions_sum_eq Z_rw Z_rw; aw.
  move=> [[fu [su tu]] _][[fv [sv tv]] _] ge.
  apply function_exten3 =>//; try ue.
move=> y; rewrite /set_of_graph_sum_eq; aw.
Qed.

Lemma sof_sum_le_equi: forall F n, inc n Bnat ->
  (set_of_functions_sum_le F n) \Eq (set_of_graph_sum_le F n).
Proof.
move=> F n nB.
exists (BL graph (set_of_functions_sum_le F n) (set_of_graph_sum_le F n)).
aw; ee; apply bl_bijective.
    move => t ts; rewrite /set_of_graph_sum_le; aw; ex_tac.
  move=> u v; rewrite /set_of_functions_sum_le Z_rw Z_rw; aw.
  move=> [[fu [su tu]] _][[fv [sv tv]] _] ge.
  apply function_exten3 =>//; try ue.
move=> y; rewrite /set_of_graph_sum_le; aw.
Qed.

Lemma set_of_functions_sum0: forall f,
  (forall a, inc a Bnat -> f 0c a = 1c) ->
  (forall a, inc a Bnat -> f a 0c = 1c) ->
  (forall a b, inc a Bnat -> inc b Bnat ->
     f (succ a) (succ b) = (f (succ a) b) +c (f a (succ b))) ->
  forall a b, inc a Bnat -> inc b Bnat -> f a b = (binom (a +c b) a).
Proof.
move=> f p2 p3 p4.
move=> a b aB bB; move: a aB b bB.
apply: cardinal_c_induction.
  move=> b bB; rewrite p2; aw; [ rewrite binom0 =>// | fprops].
move=> n nB fnb.
move=> b bB.
rewrite (plus_via_succ1 _ nB).
move: b bB; apply cardinal_c_induction.
  have snB: inc (succ n) Bnat by hprops.
  by rewrite (Bnat0_unit_sumr nB) (binom_nn snB) (p3 _ snB).
move=> c cB fsn.
have sc: inc (succ c) Bnat by hprops.
rewrite (p4 _ _ nB cB) (fnb _ sc) fsn.
by rewrite - (plus_via_succ _ cB) (binomSnSm (BS_plus nB sc) nB).
Qed.

Lemma set_of_functions_sum1: forall E x n,
  inc n Bnat -> ~ (inc x E) ->
  (set_of_graph_sum_le E n) \Eq (set_of_graph_sum_eq (tack_on E x) n).
Proof.
move=> E x n nB nxE.
set (K:= interval_cc_0a n).
set (f:= fun z=> tack_on z (J x (card_sub n (cardinal_sum z)))).
exists (BL f (set_of_graph_sum_le E n)
    (set_of_graph_sum_eq (tack_on E x) n)).
aw; ee; apply bl_bijective.
- move => z; rewrite setof_suml_aux // setof_sume_aux //.
  move=> [fgz [dz [lez alc]]].
  have p0: (fgraph (f z)) by rewrite /f; apply tack_on_fgraph => //; ue.
  have p1: (tack_on E x) = domain (f z) by rewrite -dz domain_tack_on //.
  have xd: inc x (domain (f z)) by rewrite -p1; fprops.
  have Vx2: V x (f z) = (card_sub n (cardinal_sum z)).
    rewrite /f tack_on_V_out //; ue.
  have sB: inc (cardinal_sum z) Bnat by Bnat_tac.
  move: (card_sub_pr0 nB sB lez) => [dB psb].
  ee.
    move: (partition_tack_on nxE); rewrite p1.
    move=> p2; rewrite (cardinal_sum_assoc p0 p2); bw;rewrite card_plus_pr0; bw.
    rewrite trivial_cardinal_sum3 // Vx2.
    have ->: (restr (f z) E) = z =>//.
      rewrite /restr /f; aw; set_extens t; rewrite Z_rw; aw.
        move=> [tgz tJ]; case tgz =>//.
        move=> aux; rewrite aux in tJ; awi tJ; contradiction.
      move=> tgz; ee; rewrite -dz; apply (inc_pr1_domain tgz).
      by rewrite cardinal_of_cardinal; fprops.
   rewrite -p1; move=> i idx; awi idx; case idx.
   rewrite -dz /f; move=> h; rewrite tack_on_V_in //; try ue; apply alc=>//.
  move => ->; rewrite Vx2. fprops.
- move => u v; rewrite setof_suml_aux // setof_suml_aux //.
  move=> [fu [su [tu _]]][fv [sv [tv _]]].
  apply tack_on_g_injective => //; ue.
- move=> y; rewrite setof_sume_aux //.
  move=> [fy [sy [ty suy]]].
  move: (partition_tack_on nxE) => pf; rewrite -sy in pf.
  move: (cardinal_sum_assoc fy pf).
  bw;rewrite card_plus_pr0; bw.
  have idy: inc x (domain y) by rewrite sy; fprops.
  have cvy: is_cardinal (V x y) by apply suy.
  have Ed: sub E (domain y) by rewrite sy; fprops.
  have dr: domain (restr y E) = E by rewrite restr_domain1.
  rewrite trivial_cardinal_sum3 //; move => aux; rewrite aux in ty.
  have cs: is_cardinal (cardinal_sum (restr y E)).
    by rewrite /cardinal_sum;fprops.
  exists (restr y E); ee.
    rewrite setof_suml_aux;ee; first by rewrite -ty;apply sum_increasing3;fprops.
    move=> i; rewrite restr_domain1 //; move=> iE; rewrite restr_ev //.
    apply suy; ue.
   rewrite /f.
   have ->: (n -c (cardinal_sum (restr y E)) = V x y).
   rewrite - (cardinal_of_cardinal cvy).
   rewrite -ty in nB.
   apply card_sub_pr2.
     apply (Bnat_in_sum (cardinal_cardinal (V x y)) nB).
     apply (Bnat_in_sum2 cs nB).
   rewrite card_plusC //.
  apply tack_on_restr =>//.
Qed.

Lemma set_of_functions_sum2: forall E n, inc n Bnat ->
 cardinal(set_of_graph_sum_le E (succ n))
 = (cardinal (set_of_graph_sum_eq E (succ n)))
    +c (cardinal (set_of_graph_sum_le E n)).
Proof.
move => E n nB.
have snB: inc (succ n) Bnat by hprops.
set A:= (set_of_graph_sum_eq E (succ n)).
set B:= (set_of_graph_sum_le E (succ n)).
set C:= (set_of_graph_sum_le E n).
have di: disjoint A C.
  apply disjoint_pr.
  move=> u; rewrite /A/C setof_suml_aux // setof_sume_aux //.
  move=> [fgx [dxE [sles alc]]] [_ [_ [nes _]]].
  rewrite sles in nes; move: (is_lt_succ nB) => aux; co_tac.
have ->: B = union2 A C.
  set_extens t; rewrite union2_rw /A/B/C ? setof_suml_aux // setof_sume_aux //.
  move=> [fgx [dxE [sles alc]]].
    case (equal_or_not (cardinal_sum t) (succ n)) => h; first by left; ee.
    have : cardinal_lt (cardinal_sum t) (succ n) by split => //.
    rewrite lt_is_le_succ =>//; fprops.
  case; move=> [fgx [dxE [sles alc]]]; ee.
   rewrite sles; fprops.
  apply cardinal_le_transitive with n =>//.
  apply is_le_succ; fprops.
apply card_plus_pr1 =>//; fprops.
Qed.

Lemma set_of_functions_sum3: forall E,
  cardinal (set_of_graph_sum_le E 0c) = 1c.
Proof.
move=> E.
have zb: inc 0c Bnat by fprops.
set w:= set_of_graph_sum_le _ _.
suff ->: (w = singleton (L E (fun _ =>0c))) by apply cardinal_singleton.
set_extens g.
  move=> gw; move: (gw); rewrite /w setof_suml_aux; aw.
  move=> [fg [dg [lez ac]]]; apply fgraph_exten =>//; [gprops | bw | ].
  rewrite dg; move=> x xe; bw.
  move: gw; rewrite /w/set_of_graph_sum_le; aw; move=> [z].
  rewrite /set_of_functions_sum_le Z_rw;aw.
  move=> [[[fz [sz tz]] _] gz].
  rewrite -sz in xe; move: (inc_W_target fz xe); rewrite tz.
  rewrite interval_Bnat_pr0 // /W gz=> aux.
  apply (zero_smallest2 aux).
have p: cardinal_le (cardinal_sum (cst_graph E 0c)) 0c.
  rewrite sum_of_same1 card_multC zero_prod_absorbing; fprops.
aw; move=> ->; rewrite /w setof_suml_aux //; ee; bw.
move=> i iE; bw; fprops.
Qed.

Lemma set_of_functions_sum4: forall n, inc n Bnat->
    cardinal (set_of_graph_sum_le emptyset n) = 1c.
Proof.
move=> n nB.
suff: (set_of_graph_sum_le emptyset n =singleton emptyset).
  by move=> ->;apply cardinal_singleton.
set_extens f; aw; rewrite setof_suml_aux //.
  move=> [g [df _]]; apply is_emptyset; move=> y yf.
  move: (inc_pr1_domain yf);rewrite df => h; elim (emptyset_pr h).
move=> ->; ee; srw.
   apply emptyset_fgraph.
rewrite trivial_cardinal_sum; fprops; srw.
move=> i ie; elim (emptyset_pr ie).
Qed.

Lemma set_of_functions_sum_pr: forall n h,
  inc n Bnat -> inc h Bnat ->
  let intv:= fun h => (interval_co_0a h) in
    let sle:= fun n h => set_of_graph_sum_le (intv h) n in
      let seq := fun n h => set_of_graph_sum_eq (intv h) n in
        let A:= fun n h => cardinal (sle n h) in
          let B:= fun n h => cardinal (seq n h) in
            (A n h = B n (succ h) & A n h = (binom (n +c h) n)).
Proof.
move=> n h nB hB intv sle seq A B.
have AB: forall a b, inc a Bnat -> inc b Bnat -> A a b = B a (succ b).
  move=> a b aB bB; rewrite /A /B /sle/seq; aw.
  move: (interval_co_pr4 bB) => [p1 p2]; rewrite -/intv in p1 p2.
  rewrite -p1; apply (set_of_functions_sum1 aB p2).
split; first by apply AB.
have Hv: forall a b, inc a Bnat -> inc b Bnat ->
  A (succ a) b = (B (succ a) b) +c (A a b).
  move=> a b aB bB; rewrite /A/B /sle /seq.
  by rewrite - (set_of_functions_sum2 (intv b) aB).
apply set_of_functions_sum0.
- move=> a aB; rewrite /A/sle; apply set_of_functions_sum3.
- move=> a aB;rewrite /A/sle.
  have ->: (intv 0c = emptyset).
    by rewrite /intv;apply emptyset_interval_00.
  apply set_of_functions_sum4 =>//.
- move=> a b aB bB.
  by rewrite (Hv _ _ aB (BS_succ bB)) -(AB _ _ (BS_succ aB) bB).
- exact.
- exact.
Qed.

Definition set_of_graph_sum_le_int p n :=
  set_of_graph_sum_le (interval_cc_0a p) n.

Definition set_of_increasing_functions_int p n :=
  (Zo (set_of_functions (interval_cc_0a p) (interval_cc_0a n))
    (fun z => increasing_fun z
      (interval_Bnato 0c p)
      (interval_Bnato 0c n))).

Lemma card_set_of_increasing_functions_int : forall p n,
  inc p Bnat -> inc n Bnat ->
  cardinal (set_of_increasing_functions_int p n) =
  binom (succ (n +c p)) (succ p).
Proof.
move=> p n pB nB.
have spB: inc (succ p) Bnat by hprops.
rewrite -(plus_via_succ _ pB).
rewrite - (cardinal_set_of_increasing_functions3 nB spB).
rewrite /set_of_increasing_functions_int /set_of_incr_functions.
have zB: inc 0c Bnat by fprops.
rewrite interval_Bnato_substrate // interval_Bnatco_substrate //.
rewrite -interval_co_cc //.
have->: (interval_Bnatco (succ p) = (interval_Bnato 0c p)).
  rewrite /interval_Bnato /interval_Bnatco /interval_co_0a.
  apply f_equal.
  set_extens x; rewrite Bnat_interval_co_pr1 // Bnat_interval_cc_pr1 //
     lt_is_le_succ //.
by[].
Qed.

Lemma double_restrc: forall f n p, fgraph f -> inc p Bnat ->
  n <c p ->
  domain f = interval_cc_0a p ->
  restr (restr f (interval_cc_0a (succ n)))
     (interval_cc_0a n) =
     restr f (interval_cc_0a n).
Proof.
move=> f n p fgf pB ltnp df.
set (E1:= interval_cc_0a n).
set (E2:= interval_cc_0a (succ n)).
set (E3:= interval_cc_0a p) in df.
have nB: inc n Bnat by move: ltnp=> [h _]; Bnat_tac.
have lesnp: (cardinal_le (succ n) p) by srw; fprops.
have dsnB: inc (succ n) Bnat by hprops.
have lense: cardinal_le n (succ n) by apply is_le_succ; fprops.
have s12: (sub E1 E2) by rewrite /E1 /E2; apply interval_cc_0a_increasing.
have s23: (sub E2 E3) by rewrite /E2 /E3; apply interval_cc_0a_increasing.
have s13: sub E1 E3 by apply sub_trans with E2.
have dr1: (domain (restr f E1) = E1) by rewrite restr_domain1 // df//.
have dr2: (domain (restr f E2) = E2) by rewrite restr_domain1 // df//.
have dr3: domain (restr (restr f E2) E1) = E1.
   rewrite restr_domain1 =>//;[ fprops | ue].
apply fgraph_exten; fprops.
  by rewrite dr3 dr1.
rewrite dr3; move=> x xd;rewrite ? restr_ev // ? df // .
    by apply s12.
  fprops.
by rewrite dr2.
Qed.

Lemma induction_on_sum3: forall f m,
  fgraph f -> inc m Bnat ->
  domain f = interval_cc_0a m ->
  (forall a, inc a (domain f) -> is_cardinal (V a f)) ->
  (cardinal_sum (restr f (interval_cc_0a 0c))
    = (V 0c f)
    & (forall n, n <=c m ->
      (cardinal_sum (restr f (interval_co_0a n))) +c (V n f)
      = cardinal_sum (restr f (interval_co_0a (succ n))))).
Proof.
move=> f m fgf mB df alc.
split.
  rewrite interval_zero_zero.
  have zd: (inc 0c (domain f)) by rewrite df interval_Bnat_pr0; fprops.
  rewrite trivial_cardinal_sum3 ? cardinal_of_cardinal =>//; apply alc =>//.
move=> n lenm.
have nB: inc n Bnat by Bnat_tac.
move: (partition_tack_on_intco nB) => pfa.
set (g:= restr f (interval_co_0a (succ n))).
have sid: (sub (interval_co_0a (succ n)) (domain f)).
  rewrite df -interval_co_cc //; apply interval_cc_0a_increasing => //.
have fgg: fgraph g by rewrite /g; fprops.
have dg: domain g = (interval_co_0a (succ n)).
   rewrite /g restr_domain1//.
have nd: inc n (domain f) by rewrite df interval_Bnat_pr0.
have ndg: inc n (domain g) by rewrite dg interval_co_0a_pr3; fprops.
have vfg: (V n f = V n g).
   rewrite /g restr_ev //; apply inc_a_interval_co_succ =>//.
have cv: is_cardinal (V n g) by rewrite -vfg; apply alc.
have ->: (restr f (interval_co_0a n) = restr g (interval_co_0a n)).
  rewrite /g/restr; set_extens t; rewrite Z_rw Z_rw Z_rw; last by intuition.
  by move=> [tf Pi]; ee; apply (interval_co_0a_increasing nB).
rewrite -dg in pfa.
rewrite (cardinal_sum_assoc fgg pfa); bw; rewrite card_plus_pr0; bw.
rewrite vfg trivial_cardinal_sum3 ? (cardinal_of_cardinal cv) //.
Qed.

Definition sum_to_increasing_fun y :=
  fun i => cardinal_sum (restr y (interval_cc_0a i)).

Definition sum_to_increasing_fct y n p :=
  BL (sum_to_increasing_fun y)
  (interval_cc_0a p) (interval_cc_0a n).

Lemma sum_to_increasing1: forall y n p,
  inc n Bnat -> inc p Bnat ->
  inc y (set_of_graph_sum_le_int p n) ->
  transf_axioms (sum_to_increasing_fun y)
    (interval_cc_0a p)
    (interval_cc_0a n).
Proof.
move=> y n p nB pB; rewrite /set_of_graph_sum_le_int.
rewrite setof_suml_aux //; move=> [fgy [dy [les alc]]] u.
rewrite /sum_to_increasing_fun ? interval_Bnat_pr0 =>// up.
have sj:(sub (interval_cc_0a u) (domain y)).
  by rewrite dy; apply interval_cc_0a_increasing.
move: (sum_increasing1 fgy alc sj) => aux; co_tac.
Qed.

Lemma sum_to_increasing2: forall n p,
  inc n Bnat -> inc p Bnat ->
  transf_axioms (fun y=> (sum_to_increasing_fct y n p))
  (set_of_graph_sum_le_int p n)
  (set_of_increasing_functions_int p n).
Proof.
move=> n p nB pB y ys.
move: (sum_to_increasing1 nB pB ys) => ta1.
rewrite /sum_to_increasing_fct.
have aa: (forall u, inc u (domain y)-> inc (V u y) (interval_cc_0a n)).
  move: ys;rewrite /set_of_graph_sum_le_int /set_of_graph_sum_le; aw.
  move=> [z]; rewrite/set_of_functions_sum_le Z_rw; aw.
  move=> [[[fz [sz tg]] _] hz] udy.
  rewrite -hz -tg; aw; move=> h; apply inc_W_target =>//.
move: ys;rewrite setof_suml_aux //; move=> [fgy [dy [les alc]]].
have ab: (forall u, inc u (domain y) -> inc (V u y) Bnat).
  move=> u udy; apply (sub_interval_Bnat (aa _ udy)).
set g:= BL _ _ _ .
have p1: is_function g by rewrite /g;apply bl_function.
rewrite /set_of_increasing_functions_int; apply Z_inc.
  rewrite /g;aw; ee;aw.
have zB: inc 0c Bnat by fprops.
red; ee.
- by move: (interval_Bnato_worder inc0_Bnat pB) => [ok_ ].
- by move: (interval_Bnato_worder inc0_Bnat nB) => [ok_ ].
- by rewrite interval_Bnato_substrate /g; aw.
- by rewrite interval_Bnato_substrate /g; aw.
- move=> a b; rewrite interval_Bnato_related // interval_Bnato_related //.
  move => [asg [bsg leab]].
  split.
    have ag: inc a (source g) by rewrite /g; aw.
    move: (inc_W_target p1 ag); rewrite {2} /g; aw.
  split.
    have bg: inc b (source g) by rewrite /g; aw.
    move: (inc_W_target p1 bg); rewrite {2} /g; aw.
  rewrite /g bl_W // bl_W //.
  rewrite /sum_to_increasing_fun.
  set f:= restr y (interval_cc_0a b).
  have fgf: fgraph f by rewrite /f; fprops.
  move: (asg)(bsg); rewrite ? interval_Bnat_pr0 // => lap lbp.
  have si: sub (interval_cc_0a b) (domain y).
    by rewrite dy; apply interval_cc_0a_increasing.
  have df: domain f = (interval_cc_0a b) by rewrite restr_domain1 //.
  have fc:(forall x, inc x (domain f) -> is_cardinal (V x f)).
     move=> x; rewrite df; move=> xdf; rewrite /f restr_ev //; apply alc.
     apply si=>//.
  set j := (interval_cc_0a a).
  have bB: inc b Bnat by Bnat_tac.
  have sj: sub j (domain f) by rewrite /f df;apply interval_cc_0a_increasing.
  have ->: restr y j = restr f j. rewrite /f double_restr // -df //.
  apply (sum_increasing1 fgf fc sj).
Qed.

Lemma sum_to_increasing4: forall n p,
  inc n Bnat -> inc p Bnat ->
  injection (BL (fun y=> (sum_to_increasing_fct y n p))
  (set_of_graph_sum_le_int p n)
  (set_of_increasing_functions_int p n)).
Proof.
move=> n p nB pB; apply bl_injective.
  apply sum_to_increasing2 =>//.
rewrite /sum_to_increasing_fct.
move=> u v us vs h.
move: (sum_to_increasing1 nB pB vs)=> ta2.
have aux: forall x, inc x (interval_cc_0a p) ->
  cardinal_sum (restr u (interval_cc_0a x)) =
   cardinal_sum (restr v (interval_cc_0a x)).
  move=> x xi; move: (f_equal (W x) h); rewrite ? bl_W //.
  apply sum_to_increasing1 =>//.
clear h.
move: us vs; rewrite /set_of_graph_sum_le_int ? setof_suml_aux //.
move=>[fgu [du [lesu alcu]]] [fgv [dv [lesv alcv]]].
apply fgraph_exten =>//; first by ue.
move: (induction_on_sum3 fgu pB du alcu) => [u1 u2].
move: (induction_on_sum3 fgv pB dv alcv) => [v1 v2].
move=> x xdu.
case (equal_or_not x 0c)=> zx.
  rewrite zx -u1 -v1 aux =>//; rewrite interval_Bnat_pr0 //; fprops.
have xlep: cardinal_le x p by move: xdu; rewrite du interval_Bnat_pr0 //.
have xB: inc x Bnat by rewrite du in xdu; apply (sub_interval_Bnat xdu).
move: (predc_pr xB zx) => [prB xsp].
have i2:(interval_co_0a x = (interval_cc_0a (predc x))).
  rewrite {1} xsp -interval_co_cc //.
have xip: inc x (interval_cc_0a p) by ue.
have xpip: inc (predc x) (interval_cc_0a p).
have psp: cardinal_le p (succ p) by apply is_le_succ; fprops.
   rewrite interval_Bnat_pr0 //; rewrite lt_n_succ_le1 //; last by fprops.
   rewrite -xsp; co_tac.
move: (u2 _ xlep)(v2 _ xlep); rewrite -interval_co_cc // i2.
move: (ta2 _ xip); rewrite / sum_to_increasing_fun.
rewrite (aux _ xip) (aux _ xpip).
set A:= cardinal_sum _; set B:= cardinal_sum _.
move=> AB h1 h2; move: (sub_interval_Bnat AB) => AB1.
apply plus_simplifiable_left with B => //.
  have Bc: is_cardinal B by rewrite /B /cardinal_sum; fprops.
   rewrite -h1 in AB1;apply (Bnat_in_sum2 Bc AB1).
  rewrite -h1 in AB1; apply (Bnat_in_sum (alcu _ xdu) AB1).
  rewrite du -dv in xdu.
  rewrite -h2 in AB1; apply (Bnat_in_sum (alcv _ xdu) AB1).
rewrite h1 h2//.
Qed.

Lemma sum_to_increasing5: forall n p,
  inc n Bnat -> inc p Bnat ->
  surjection (BL (fun y=> (sum_to_increasing_fct y n p))
  (set_of_graph_sum_le_int p n)
  (set_of_increasing_functions_int p n)).
Proof.
move=> n p nB pB; apply bl_surjective.
  apply sum_to_increasing2 =>//.
move =>y; rewrite /set_of_increasing_functions_int Z_rw; aw.
move=> [[fy [sy ty]] iy].
rewrite /set_of_graph_sum_le_int /sum_to_increasing_fct.
set E1:= interval_cc_0a p in sy iy |- *.
set E2:= interval_cc_0a n in ty iy |- *.
have Hb:inc 0c E1. rewrite /E1 interval_Bnat_pr0; fprops.
set (a:= W 0c y).
have aE2: inc a E2 by rewrite -ty /a; apply inc_W_target =>//; ue.
have rec:(forall u, u<> 0c -> inc u (source y) ->
    exists v, inc v (source y) & u = succ v & v = predc u
      & (W v y) <=c (W u y)).
  move=> u unz usy.
  have up: cardinal_le u p by move: usy;rewrite sy interval_Bnat_pr0 //.
  have uB:inc u Bnat by Bnat_tac.
  move: (predc_pr uB unz)=> [pn sp].
  have cp: is_cardinal (predc u) by fprops.
  move: (is_le_succ0 cp); rewrite -sp => lepu.
  have psy: (inc (predc u) (source y)).
   rewrite sy interval_Bnat_pr0 //; co_tac.
  have zB: inc 0c Bnat by fprops.
  ex_tac;ee; move: iy => [_ [o1 [o2 [s1 [s2 op]]]]].
  have ge1: gle (interval_Bnato 0c p) (predc u) u.
    rewrite /E1 /interval_cc_0a in sy.
    rewrite interval_Bnato_related // -sy;ee.
  move: (op _ _ ge1); rewrite interval_Bnato_related //; intuition.
set (f:= fun i => Yo(i= 0c) a (card_sub (W i y) (W (predc i) y))).
have f1p: (forall i, inc i E1 -> inc (f i) E2).
  move=> i i1; rewrite /f; Ytac h =>//.
  rewrite sy in rec; move: (rec _ h i1) => [v [vsy [usv [vp wle]]]].
  rewrite /E2 interval_Bnat_pr0 //.
  rewrite -sy in i1 vsy.
  move : (inc_W_target fy i1)(inc_W_target fy vsy).
  rewrite ty interval_Bnat_pr0 // interval_Bnat_pr0 //; move=> win wvn.
  have wib: inc (W i y) Bnat by Bnat_tac.
  have wvb: inc (W v y) Bnat by Bnat_tac.
  move: (sub_le_symmetry wle) => h1; rewrite -vp; co_tac.
set (g:=L E1 f).
have fg: fgraph g by rewrite /g; gprops.
have dg: domain g = interval_cc_0a p by rewrite /g;bw.
have cg: forall a, inc a (domain g) -> is_cardinal (V a g).
  rewrite /g; bw;move=> b bdg; rewrite /g; bw.
  move: (f1p _ bdg);rewrite /E2 => aux; move: (sub_interval_Bnat aux) => wB.
  fprops.
move: (induction_on_sum3 fg pB dg cg)=> [g0 g1].
set (h:= fun i=> cardinal_sum (restr g (interval_cc_0a i))).
have h0: (h 0c = W 0c y).
  rewrite /h g0 /g; bw; rewrite /f Y_if_rw //.
have h1: forall i, inc i E1 -> h i = W i y.
  move=> i; rewrite /E1 interval_Bnat_pr0 // => ip.
  apply (cardinal_c_induction5 (r:=fun i=> h i = W i y) pB h0) =>//.
  move=> m mp; rewrite /h => hr.
  have mB: inc m Bnat by move: mp => [mp _]; Bnat_tac.
  have cm: is_cardinal m by fprops.
  have smB: inc (succ m) Bnat by hprops.
  set (n1:= succ m).
  have n1p: (cardinal_le n1 p) by rewrite /n1; srw; fprops.
  move: (g1 _ n1p); rewrite /n1.
  rewrite - interval_co_cc // -interval_co_cc // hr; move=> <-.
  have snz: (succ m <> 0c) by apply succ_nonzero.
  have sm1: inc (succ m) E1 by rewrite /E1 interval_Bnat_pr0 //; srw; fprops.
  rewrite /g; bw; rewrite /f Y_if_not_rw // predc_pr1 //.
  have smsy: (inc (succ m) (source y)) by ue.
  move: (rec _ snz smsy) => [v [vsy [smsv [vp wle]]]].
  have mv: m = v.
    apply succ_injective1; fprops.
    move: vsy; rewrite sy interval_Bnat_pr0 //; move => [aux _]; fprops.
  rewrite -mv in wle.
  move: (inc_W_target fy smsy); rewrite ty /E2 => wi.
  move: (sub_interval_Bnat wi)=> wsB.
  apply card_sub_pr =>//; Bnat_tac.
have p1: inc g (set_of_graph_sum_le E1 n).
  rewrite /set_of_graph_sum_le setof_suml_aux //.
  ee.
  have pe1: inc p E1 by rewrite /E1 interval_Bnat_pr0//; fprops.
  move: (h1 _ pe1); rewrite /h -dg restr_to_domain; fprops.
  move => ->.
  rewrite -sy in pe1; move: (inc_W_target fy pe1).
  rewrite ty /E2 interval_Bnat_pr0//.
ex_tac.
move :(sum_to_increasing1 nB pB p1) => aux.
symmetry; apply function_exten => //; aw.
  apply bl_function => //.
rewrite sy; move=> x xsy; rewrite bl_W //.
rewrite -h1 //.
Qed.

Lemma sum_to_increasing6: forall n p,
  inc p Bnat -> inc n Bnat ->
  cardinal (set_of_graph_sum_le_int p n) =
  binom (succ (n +c p)) (succ p).
Proof.
move=> n p pB nB.
rewrite - card_set_of_increasing_functions_int //.
aw.
 exists (BL (fun y=> (sum_to_increasing_fct y n p))
    (set_of_graph_sum_le_int p n)
    (set_of_increasing_functions_int p n)).
aw; ee.
by split; [apply sum_to_increasing4 | apply sum_to_increasing5].
Qed.

End IntegerProps.
Export IntegerProps.