Library sset3
Theory of Sets: EII-4 Union and intersection of a family of sets
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 sset2.
Module Bunion.
Hint Resolve fcompose_fgraph inc_V_range: fprops.
Bourbaki defines union and intersection of a functional graph
Record Uintegral (I :Set)(f :I->Set) : Set := {UI_z : I; UI_elt : f UI_z}.
Definition uniont (I:Set)(f : I->Set) :=
IM (fun i : Uintegral f => Ro (UI_elt i)).
Definition intersectiont (I:Set) (f : I->Set):=
Zo (uniont f) (fun y => forall z : I, inc y (f z)).
Definition unionf (x:Set)(f: Set->Set) := uniont (fun a:x => f (Ro a)).
Definition unionb (g:Set) := unionf (domain g)(fun a=> V a g).
Definition intersectionf (x:Set)(f: Set->Set):=
intersectiont(fun a:x => f (Ro a)).
Definition intersectionb (g:Set) := intersectionf (domain g) (fun a=> V a g).
Basic properties of union an intersection
Lemma uniont_rw: forall (I:Set) (f:I->Set) x,
inc x (uniont f) = exists z, inc x (f z).
Proof.
rewrite /uniont=> I f x; rewrite IM_rw; apply iff_eq.
by move=>[a] <-; exists (UI_z a); apply R_inc.
by move => [z xf]; exists (Build_Uintegral (Bo xf)); rewrite B_eq.
Qed.
Lemma unionf_rw: forall x i f,
inc x (unionf i f) = exists y, inc y i & inc x (f y).
Proof.
move=> x i f;rewrite /unionf uniont_rw ; apply iff_eq.
by move=> [z zp]; exists (Ro z); split=>//; apply R_inc.
by move=>[y [yi xf]]; exists (Bo yi); rewrite B_eq.
Qed.
Lemma unionb_rw: forall x f,
inc x (unionb f) = exists y, inc y (domain f) & inc x (V y f).
Proof. by move=> x f ;rewrite /unionb unionf_rw. Qed.
Hint Rewrite uniont_rw unionf_rw unionb_rw: sw.
Lemma uniont_inc : forall (I:Set) (f : I->Set) y x,
inc x (f y) -> inc x (uniont f).
Proof. by move=> I f y x; srw; exists y. Qed.
Lemma unionf_inc: forall x y i f,
inc y i -> inc x (f y) -> inc x (unionf i f).
Proof. by move=> x y i f yi xf;srw; exists y. Qed.
Lemma unionf_exists: forall x i f,
inc x (unionf i f) -> exists y, inc y i & inc x (f y).
Proof. by move=> x i f; srw. Qed.
Lemma unionb_inc: forall x y f,
inc y (domain f) -> inc x (V y f) ->
inc x (unionb f).
Proof. by move=> x y f yd xv; srw; ex_tac. Qed.
Lemma unionb_exists: forall x f,
inc x (unionb f) ->
exists y, inc y (domain f) & inc x (V y f).
Proof. by move=> x f; srw. Qed.
Ltac union_tac:=
match goal with
| H:inc ?x (?f ?y) |- inc ?x (uniont ?f)
=> apply (uniont_inc H)
| Ha : inc ?i (domain ?g), Hb : inc ?x (V ?i ?g) |- inc ?x (unionb ?g)
=> apply (unionb_inc Ha Hb)
| Ha : inc ?x ?y, Hb : inc ?y ?a |- inc ?x (union ?a)
=> apply (union_inc Ha Hb)
| Ha : inc ?y ?i, Hb : inc ?x (?f ?y) |- inc ?x (unionf ?i ?f)
=> apply (unionf_inc _ Ha Hb)
| Ha : inc ?y ?i |- inc ?x (unionf ?i ?f)
=> apply unionf_inc with y; fprops
| Ha : inc ?i (domain ?g) |- inc ?x (unionb ?g)
=> apply unionb_inc with i; fprops
| Ha : inc ?y ?i |- inc ?x (unionf ?i ?f)
=> apply unionf_inc with y; fprops
| Ha : inc ?x (V ?i ?g) |- inc ?x (unionb ?g)
=> apply unionb_inc with i; fprops
| Hb : inc ?x (?f ?y) |- inc ?x (unionf ?i ?f)
=> apply unionf_inc with y; fprops
end.
Lemma intersectiont_rw: forall (I:Set) (f:I-> Set) x, nonempty I ->
inc x (intersectiont f) = (forall j, inc x (f j)).
Proof.
rewrite/intersectiont=> I f x [j jI]; apply iff_eq; rewrite Z_rw.
intuition.
move=> hyp; split=>//; srw; case jI=> [t tI]; exists t; apply: hyp.
Qed.
Lemma intersectionf_rw : forall I f x, nonempty I ->
inc x (intersectionf I f) = (forall j, inc j I -> inc x (f j)).
Proof.
move=> I f x [y yI].
rewrite /intersectionf intersectiont_rw; last by exists y.
apply iff_eq.
by move=> hyp j jI; move: (hyp (Bo jI)); rewrite B_eq.
by move=> hyp j; apply hyp; apply R_inc.
Qed.
Lemma nonempty_domain: forall x,
nonempty (domain x) = nonempty x.
Proof.
rewrite/domain =>x;apply iff_eq; move=> [y].
by aw; move=>[z [zx]]; exists z.
by move=> yx; exists (P y); aw; ex_tac.
Qed.
Lemma intersectionb_rw : forall g x, nonempty g ->
inc x (intersectionb g) = (forall i, inc i (domain g) -> inc x (V i g)).
Proof.
move=> g x neg; rewrite /intersectionb intersectionf_rw //.
by rewrite nonempty_domain.
Qed.
Lemma intersection_rw : forall x y, nonempty y ->
inc x (intersection y) = (forall i, inc i y -> inc x i).
Proof.
move=> x y ney; apply iff_eq;first by move=> xi iy; apply intersection_forall.
by apply intersection_inc.
Qed.
Hint Rewrite intersectiont_rw intersectionf_rw intersectionb_rw intersection_rw : sw.
Lemma intersectiont_inc : forall (I:Set) (f:I-> Set) x, nonempty I ->
(forall j, inc x (f j)) -> inc x (intersectiont f).
Proof. move=> I f x nI; srw. Qed.
Lemma nonempty_aux: forall (I:Set) (i:I), nonempty I.
Proof. move=> I i;exists (Ro i); apply R_inc. Qed.
Lemma intersectiont_forall : forall (I:Set) (f:I-> Set) x j,
inc x (intersectiont f) -> inc x (f j).
Proof. move=> I f x j; srw. apply: (nonempty_aux j). Qed.
Lemma intersectionf_inc :forall I f x, nonempty I ->
(forall j, inc j I -> inc x (f j)) -> inc x (intersectionf I f).
Proof. move => I f x nI; srw. Qed.
Lemma intersectionf_forall :forall I f x j,
inc x (intersectionf I f) -> inc j I -> inc x (f j).
Proof. move=> I f x j hyp jI; move: hyp; srw; [ by apply | ex_tac]. Qed.
Lemma intersectionb_inc : forall g x, nonempty g ->
(forall i, inc i (domain g) -> inc x (V i g)) -> inc x (intersectionb g).
Proof. move=> g x neg; srw. Qed.
Lemma intersectionb_forall : forall g x i,
inc x (intersectionb g) -> inc i (domain g) -> inc x (V i g).
Proof.
move=> g x i xi id; move: xi; srw; first by apply.
by rewrite -nonempty_domain; ex_tac.
Qed.
Lemma intersectiont_empty: forall (I:Set) (f:I-> Set),
I = emptyset -> intersectiont f = emptyset.
Proof.
move=> I f nneI; empty_tac t Ht.
move: Ht; rewrite /intersectiont Z_rw; srw.
move => [[z _] _];rewrite nneI in z; case z.
Qed.
Lemma intersectionf_empty: forall f,
intersectionf emptyset f = emptyset.
Proof.
by rewrite /intersectionf => f; apply intersectiont_empty.
Qed.
Lemma intersectionb_empty:
intersectionb emptyset = emptyset.
Proof.
rewrite /intersectionb; srw; apply intersectionf_empty.
Qed.
Lemma uniont_exten: forall (I:Set) (f: I-> Set) (f':I->Set),
(forall i, f i = f' i) -> uniont f = uniont f'.
Proof.
move=> I f f' hyp; set_extens t; srw; move=> [x xp]; exists x.
by rewrite -hyp.
by rewrite hyp.
Qed.
Lemma unionf_exten: forall sf f f',
(forall i, inc i sf -> f i = f' i) -> unionf sf f = unionf sf f'.
Proof.
move=> sf f f' hyp; set_extens t; srw; move=> [x [xp1 xp2]]; ex_tac.
by rewrite -hyp.
by rewrite hyp.
Qed.
Lemma unionb_exten: forall f f',
f = f' -> unionb f = unionb f'.
Proof. by move=> f f' ->.
Qed.
Lemma intersectiont_exten:
forall (I:Set) (f:I-> Set) (f':I-> Set),
(forall i, f i = f' i) -> (intersectiont f) = (intersectiont f').
Proof.
move=> I f f' hyp; case (emptyset_dichot I) => ie.
do 2 rewrite intersectiont_empty //.
by set_extens t; srw; move=> jp j; move: (jp j); rewrite hyp.
Qed.
Lemma intersectionf_exten: forall I f f',
(forall i, inc i I -> f i = f' i) ->
intersectionf I f = intersectionf I f'.
Proof.
move => I f f' hyp; case (emptyset_dichot I)=> pI.
by rewrite pI; do 2 rewrite intersectionf_empty.
set_extens t; srw; move=> jp j jI; move: (jp _ jI); rewrite hyp //.
Qed.
Lemma intersectionb_exten: forall g g',
g = g' -> intersectionb g = intersectionb g'.
Proof. by move=> g g'; apply f_equal. Qed.
Lemma uniont_sub: forall (I:Set) (f: I-> Set) i,
sub (f i) (uniont f).
Proof. move=> I f i t tf; union_tac. Qed.
Lemma intersectiont_sub: forall (I:Set) (f: I-> Set) i,
sub (intersectiont f) (f i).
Proof. move=> I f i t. apply (intersectiont_forall). Qed.
Lemma sub_uniont: forall (I:Set) (f: I-> Set) x,
(forall i, sub (f i) x) -> sub (uniont f) x.
Proof. by move=> I f x hyp t; srw; move=> [z tfz]; apply (hyp z). Qed.
Lemma sub_intersectiont: forall (I:Set) (f: I-> Set) x,
nonempty I ->
(forall i, sub x (f i)) -> sub x (intersectiont f).
Proof. by move=> I f x ne hyp t tx; srw => j; apply hyp. Qed.
Lemma intersectiont_sub2: forall (I:Set) (f: I-> Set) x,
(forall i, sub (f i) x) -> sub (intersectiont f) x.
Proof.
move=> I f x hyp t ti.
case (emptyset_dichot I) => ie.
move: ti; rewrite intersectiont_empty // => te; empty_tac0.
case ie => y; case => z _; apply (hyp z); move: ti;srw.
Qed.
Lemma sub_uniont2: forall (I:Set) (f: I-> Set) x,
nonempty I -> (forall i, sub x (f i)) -> sub x (uniont f).
Proof. by move=> I f x [y yI] hyp t tx; srw; exists (Bo yI);apply: hyp. Qed.
Lemma empty_uniont1: forall (I:Set) (f: I-> Set),
nonempty (uniont f) -> nonempty I.
Proof. move=> I f [x]; srw; move => [z _]; apply (nonempty_aux z). Qed.
Lemma empty_unionf1: forall sf f,
nonempty (unionf sf f) -> nonempty sf.
Proof. by move=> sf f [x]; srw; move => [y [ ysf _]]; ex_tac. Qed.
Lemma empty_unionf: forall sf f,
sf = emptyset -> unionf sf f = emptyset.
Proof.
move=> s f ->; empty_tac t ht; move: ht; srw.
move=> [y [ye _]]; empty_tac0.
Qed.
Proposition 1 : union (compose g f)
Theorem uniont_rewrite: forall (I K:Set) (f: K->I) (g:I ->Set),
surjectiveC f ->
uniont g = uniont (fun k:K => g(f k)).
Proof.
move=> I K f g sf; set_extens t; srw; move=> [z zp]; last by exists (f z).
by move: (sf z) => [v fv]; exists v; ue.
Qed.
Theorem intersectiont_rewrite: forall (I K:Set) (f: K->I) (g:I ->Set),
surjectiveC f ->
intersectiont g = intersectiont (fun k:K => g(f k)).
Proof.
move=> I K f g sf; case (emptyset_dichot I) => ei.
have ek: K = emptyset.
apply is_emptyset=> y yK ;case yK => x _; move: (f x); rewrite ei; case.
rewrite ! intersectiont_empty=>//.
have nek: nonempty K.
move: ei => [k [i]] _; move: (sf i)=> [j _]; apply (nonempty_aux j).
set_extens t; srw; move=> hyp j; move: (sf j)=> [i] <-; apply hyp.
Qed.
Lemma unionb_rewrite1: forall f g,
is_function f -> fgraph g -> range (graph f) = domain g ->
unionb g = unionb (fcompose g (graph f)).
Proof.
move=> f g [_ [ fgf _]] fgg rfg.
have fc: fcomposable g (graph f) by hnf; intuition; ue.
set_extens t; srw; rewrite (fcomposable_domain fc) -rfg.
move=> [y [yr tV]]; move: yr; srw; move=> [x [xd yv]]; ex_tac; bw; ue.
by move=> [y [yd]]; bw=> tV; exists (V y (graph f)); split =>//; fprops.
Qed.
Lemma unionb_rewrite: forall f g,
fgraph f -> range f = domain g ->
unionb g = unionb (gcompose g f).
Proof.
rewrite /gcompose=> f g fgf rfsg.
set_extens t; srw; bw; rewrite - rfsg.
by move=> [y [yd tV]]; move: yd; srw; move=> [x [xd yV]]; ex_tac; bw; ue.
by move=> [y [yd]]; bw=> tV; exists (V y f); split; fprops.
Qed.
Lemma intersectionb_rewrite: forall f g,
fgraph f -> range f = domain g ->
intersectionb g = intersectionb (gcompose g f).
Proof.
move=> f g fgf rfdg.
have ec: g=emptyset -> (gcompose g f = emptyset).
move=> ge; empty_tac t tp; move: tp; rewrite /gcompose; aw;
move=> [z [zd _]].
have :(inc (V z f) (range f)) by fprops.
by rewrite rfdg ge emptyset_domain; apply emptyset_pr.
have nec: nonempty g -> (nonempty (gcompose g f)).
move=> [x xg]; rewrite /gcompose.
have p1: (inc (P x) (range f)) by rewrite rfdg /domain; dw; aw;ex_tac.
move:p1; dw; fprops; move=> [y Jf].
move: (inc_pr1_domain Jf); rewrite pr1_pair => yd.
by exists (J y (V (V y f) g)); aw; ex_tac.
case (emptyset_dichot g) => gp; first by rewrite ec gp.
move: (nec gp) => aux.
set_extens t; srw; rewrite /gcompose -rfdg; bw=> hr i id.
by bw; apply hr; fprops.
by move: id; srw; move=> [x [xd]] ->; move:(hr _ xd); bw.
Qed.
surjectiveC f ->
uniont g = uniont (fun k:K => g(f k)).
Proof.
move=> I K f g sf; set_extens t; srw; move=> [z zp]; last by exists (f z).
by move: (sf z) => [v fv]; exists v; ue.
Qed.
Theorem intersectiont_rewrite: forall (I K:Set) (f: K->I) (g:I ->Set),
surjectiveC f ->
intersectiont g = intersectiont (fun k:K => g(f k)).
Proof.
move=> I K f g sf; case (emptyset_dichot I) => ei.
have ek: K = emptyset.
apply is_emptyset=> y yK ;case yK => x _; move: (f x); rewrite ei; case.
rewrite ! intersectiont_empty=>//.
have nek: nonempty K.
move: ei => [k [i]] _; move: (sf i)=> [j _]; apply (nonempty_aux j).
set_extens t; srw; move=> hyp j; move: (sf j)=> [i] <-; apply hyp.
Qed.
Lemma unionb_rewrite1: forall f g,
is_function f -> fgraph g -> range (graph f) = domain g ->
unionb g = unionb (fcompose g (graph f)).
Proof.
move=> f g [_ [ fgf _]] fgg rfg.
have fc: fcomposable g (graph f) by hnf; intuition; ue.
set_extens t; srw; rewrite (fcomposable_domain fc) -rfg.
move=> [y [yr tV]]; move: yr; srw; move=> [x [xd yv]]; ex_tac; bw; ue.
by move=> [y [yd]]; bw=> tV; exists (V y (graph f)); split =>//; fprops.
Qed.
Lemma unionb_rewrite: forall f g,
fgraph f -> range f = domain g ->
unionb g = unionb (gcompose g f).
Proof.
rewrite /gcompose=> f g fgf rfsg.
set_extens t; srw; bw; rewrite - rfsg.
by move=> [y [yd tV]]; move: yd; srw; move=> [x [xd yV]]; ex_tac; bw; ue.
by move=> [y [yd]]; bw=> tV; exists (V y f); split; fprops.
Qed.
Lemma intersectionb_rewrite: forall f g,
fgraph f -> range f = domain g ->
intersectionb g = intersectionb (gcompose g f).
Proof.
move=> f g fgf rfdg.
have ec: g=emptyset -> (gcompose g f = emptyset).
move=> ge; empty_tac t tp; move: tp; rewrite /gcompose; aw;
move=> [z [zd _]].
have :(inc (V z f) (range f)) by fprops.
by rewrite rfdg ge emptyset_domain; apply emptyset_pr.
have nec: nonempty g -> (nonempty (gcompose g f)).
move=> [x xg]; rewrite /gcompose.
have p1: (inc (P x) (range f)) by rewrite rfdg /domain; dw; aw;ex_tac.
move:p1; dw; fprops; move=> [y Jf].
move: (inc_pr1_domain Jf); rewrite pr1_pair => yd.
by exists (J y (V (V y f) g)); aw; ex_tac.
case (emptyset_dichot g) => gp; first by rewrite ec gp.
move: (nec gp) => aux.
set_extens t; srw; rewrite /gcompose -rfdg; bw=> hr i id.
by bw; apply hr; fprops.
by move: id; srw; move=> [x [xd]] ->; move:(hr _ xd); bw.
Qed.
Union and intersection of a function with singleton range
Lemma uniont_constant: forall (I:Set) (f:I ->Set) (x:I),
is_constant_functionC f -> uniont f = f x.
Proof.
move=> I f x cf.
set_extens t; last by move=> tp;union_tac.
by srw; move => [z]; rewrite (cf x z).
Qed.
Lemma intersectiont_constant: forall (I:Set) (f:I ->Set) (x:I),
is_constant_functionC f -> intersectiont f = f x.
Proof.
move => I f x cf;move: (nonempty_aux x) => neI;set_extens t; srw.
by move => h j; rewrite (cf j x).
Qed.
Lemma singleton_type_inj: forall x (y:singleton x)(z:singleton x),
y = z.
Proof.
move=> x y z; apply R_inj.
have : (inc (Ro y) (singleton x)) by apply R_inc.
have : (inc (Ro z) (singleton x)) by apply R_inc.
by aw; move => -> -> .
Qed.
Lemma uniont_singleton:forall (a:Set) (x:a) (f: singleton (Ro x) -> Set),
uniont f = f (Bo (singleton_inc (Ro x))).
Proof.
move=>a x f; set_extens t => pt; last by union_tac.
move: pt; srw => [] [y tfy].
have <-: (y = (Bo (singleton_inc (Ro x)))) by apply singleton_type_inj.
done.
Qed.
Lemma intersectiont_singleton:forall (a:Set)(x:a) (f: singleton (Ro x) -> Set),
intersectiont f = f (Bo (singleton_inc (Ro x))).
Proof.
move=> a x f; set (j:= Bo (singleton_inc (Ro x))).
move:(nonempty_aux j) => neI; set_extens t; srw.
by move=> tf k; rewrite -(singleton_type_inj j k).
Qed.
Lemma unionf_singleton:forall f x,
unionf (singleton x) f = f x.
Proof.
move=> f x; set_extens t => tp; last by union_tac.
move: (unionf_exists tp)=> [y [aux ]]; awi aux.
ue.
Qed.
Lemma intersectionf_singleton: forall f x,
intersectionf (singleton x) f = f x.
Proof.
move=> f x; move: (nonempty_singleton x) => nx.
set_extens t; srw; first by move => h; move: (h _ (singleton_inc x)); aw.
by move=> tf j; aw => ->.
Qed.
Lemma constant_function_pr2:
forall x h, inc x (source h) -> is_constant_function h ->
exists g, exists f,
(g \coP f & h = g \co f & surjection f & is_singleton (target f)).
Proof.
move=> x h xs [fh p]; set (y := W x h).
have ssyth: sub (singleton y) (target h).
by move=> t; aw;move=> ->; rewrite /y; fprops.
set (g:= canonical_injection (singleton y) (target h)); exists g.
have ysy: inc y (singleton y) by fprops.
set (f:= constant_function (source h) (singleton y) y ysy); exists f.
have st: (is_singleton (target f)) by rewrite /f constant_target; exists y.
have ff: is_function f by apply constant_function_fun.
have tf: target f = (singleton y) by rewrite constant_target.
have sf: source f = source h by rewrite constant_source.
have sj: surjection f.
apply surjective_pr6=>//; rewrite /f tf; move =>z; aw; move=> ->.
by exists x; rewrite sf constant_W.
have c: g \coP f.
hnf; rewrite /f /g; intuition; fprops.
by rewrite tf /canonical_injection; aw.
intuition; apply function_exten;aw; fprops; try fct_tac.
by rewrite /g /canonical_injection; aw.
move=> z zh; aw; last (by ue); rewrite (p _ _ zh xs) /f constant_W // ci_W //.
Qed.
Lemma constant_function_pr3: forall (a:Set) (h:a->Set) (x:a),
is_constant_functionC h ->
exists f: a->singleton (Ro x), exists g:singleton (Ro x)->Set,
(forall u:a, h u = g (f u)) & (g (Bo (singleton_inc (Ro x))) = h x).
Proof.
move=> a h x ch.
exists (fun u=> (Bo (singleton_inc (Ro x)))).
exists (fun t:singleton (Ro x) => h x).
intuition.
Qed.
Link between union and unionb
Lemma union_prop: forall x, union x = unionf x (fun u => u).
Proof.
move=> x; set_extens t; srw; move=> [y [ty yx]]; ex_tac.
Qed.
Lemma intersection_prop: forall x, nonempty x ->
intersection x = intersectionf x (fun u => u).
Proof. move=> x nex; set_extens t; srw. Qed.
Lemma unionb_alt: forall f, fgraph f -> unionb f = union (range f).
Proof.
move=> f fgf; rewrite union_prop.
set_extens t; srw; move=> [y [ty yx]].
by exists (V y f); split; fprops.
by move: ty; srw; move=> [x [xd eq]]; ex_tac; ue.
Qed.
Lemma unionb_identity: forall x, unionb (identity_g x) = union x.
Proof.
move=> x; rewrite /unionb identity_domain union_prop.
by apply unionf_exten; move=> i ix; rewrite identity_ev.
Qed.
Monotony properties
Lemma union_monotone: forall (I:Set) (f g:I->Set),
(forall i, sub (f i) (g i)) -> sub (uniont f) (uniont g).
Proof.
by move=> I f g su t; srw; move=> [z tf]; exists z; apply su.
Qed.
Lemma intersection_monotone: forall (I:Set) (f g:I->Set),
(forall i, sub (f i) (g i)) -> sub (intersectiont f) (intersectiont g).
Proof.
move=> I f g su.
case (emptyset_dichot I) => i2; first by rewrite !intersectiont_empty //.
by move=> t; srw; move => hyp j; apply su; apply hyp.
Qed.
Lemma union_monotone2: forall A B f,
sub A B -> sub (unionf A f) (unionf B f).
Proof.
move=> A B f sAB t; srw; move=> [y [yA tfy]]; exists y; ee.
Qed.
Lemma intersection_monotone2: forall A B f,
sub A B -> nonempty A ->
sub (intersectionf B f) (intersectionf A f).
Proof.
move=> A B f sAB neA.
have neB: (nonempty B) by move:neA=> [x]; exists x; apply sAB.
by move=> x; srw; move=> p j jA; apply p; apply sAB.
Qed.
(forall i, sub (f i) (g i)) -> sub (uniont f) (uniont g).
Proof.
by move=> I f g su t; srw; move=> [z tf]; exists z; apply su.
Qed.
Lemma intersection_monotone: forall (I:Set) (f g:I->Set),
(forall i, sub (f i) (g i)) -> sub (intersectiont f) (intersectiont g).
Proof.
move=> I f g su.
case (emptyset_dichot I) => i2; first by rewrite !intersectiont_empty //.
by move=> t; srw; move => hyp j; apply su; apply hyp.
Qed.
Lemma union_monotone2: forall A B f,
sub A B -> sub (unionf A f) (unionf B f).
Proof.
move=> A B f sAB t; srw; move=> [y [yA tfy]]; exists y; ee.
Qed.
Lemma intersection_monotone2: forall A B f,
sub A B -> nonempty A ->
sub (intersectionf B f) (intersectionf A f).
Proof.
move=> A B f sAB neA.
have neB: (nonempty B) by move:neA=> [x]; exists x; apply sAB.
by move=> x; srw; move=> p j jA; apply p; apply sAB.
Qed.
Associativity of union and intersection
Theorem union_assoc: forall sf sg f g,
sf = unionf sg g ->
unionf sf f = unionf sg (fun l => unionf (g l) f).
Proof.
move=> sf sg f g ->.
set_extens t; srw; move=> [y [yu tfy]]; [ move: yu | move: tfy]; srw;
move=> [z [zsg ygz]]; exists z; split=>//; union_tac.
Qed.
Theorem intersection_assoc: forall sf sg f g,
(forall i, inc i sg -> nonempty (g i)) ->
sf = unionf sg g ->
intersectionf sf f = intersectionf sg (fun l => intersectionf (g l) f).
Proof.
move=> sf sg f g neall ->.
case (emptyset_dichot sg) => sgp.
by rewrite sgp empty_unionf // 2! intersectionf_empty.
have neu: (nonempty (unionf sg g)).
by move: sgp=> [x xsg]; move: (neall _ xsg)=> [y iy]; exists y; union_tac.
set_extens t; srw; move=> hyp j jsg.
by srw; [move => k kg; apply hyp; union_tac | by apply neall].
move: jsg; srw; move=> [y [ysg jgy]].
by apply (intersectionf_forall (hyp _ ysg) jgy).
Qed.
Theorem image_of_union: forall (I:Set) (f:I->Set) g,
is_correspondence g ->
image_by_fun g (uniont f) =
uniont (fun i => image_by_fun g (f i)).
Proof.
move=> I f g cg.
have Ha: is_graph (graph g) by fprops.
rewrite /image_by_fun.
set_extens t.
aw; move => [x [xu Jg]]; move: xu; srw; move=>[y xfy].
by exists y; aw; ex_tac.
by srw; move=>[z]; aw; move=> [x [xf Jg]]; ex_tac; union_tac.
Qed.
Theorem image_of_intersection: forall (I:Set) (f:I->Set) g,
is_correspondence g ->
sub (image_by_fun g (intersectiont f))
(intersectiont (fun i => image_by_fun g (f i))).
Proof.
rewrite /image_by_fun=> I f g cg t; aw; move=> [x].
case (emptyset_dichot I) => neI.
rewrite intersectiont_empty //; move=> [xE _]; empty_tac0.
srw; move=> [axfj Jg] j; aw; ex_tac.
Qed.
Theorem inv_image_of_intersection: forall (I:Set) (f:I->Set) g,
is_function g ->
(inv_image_by_fun g (intersectiont f)) =
(intersectiont (fun i => inv_image_by_fun g (f i))).
Proof.
move=> I f g fg.
have f1f2: forall i : I,
(fun i => inv_image_by_graph (graph g) (f i)) i=
(fun i => image_by_fun (inverse_fun g) (f i)) i.
by move=> i; rewrite -inv_image_by_fun_pr.
apply extensionality.
rewrite inv_image_by_fun_pr (intersectiont_exten f1f2).
by apply image_of_intersection; apply inverse_correspondence; fct_tac.
case (emptyset_dichot I) => ie.
by rewrite intersectiont_empty=>//; apply emptyset_sub_any.
rewrite /inv_image_by_fun.
move=> t; srw; move=> aj; move: (ie)=> [x [z zx]]; move: (aj z); aw.
move=> [u [ufx Jg]]; exists u; split=>//; srw.
move=> j; move: (aj j); aw; move=> [v [vfj Jgj]].
by rewrite - (W_pr fg Jg); rewrite (W_pr fg Jgj).
Qed.
Lemma inj_image_of_intersection: forall (I:Set) (f:I->Set) g,
injection g ->
(image_by_fun g (intersectiont f))
= (intersectiont (fun i => image_by_fun g (f i))).
Proof.
move=> I f g ig.
set(h := restriction2 g (source g) (range (graph g))).
set (i:= canonical_injection (range (graph g)) (target g)) .
have fg: (is_function g) by fct_tac.
have Hw: intersection2 (graph g) (product (source g) (target g)) = graph g.
by rewrite -intersection2_sub; case fg; case; intuition.
move: (canonical_decomposition1 fg (refl_equal h) (refl_equal i))
=> [cig [gcig [ii [sh p]]]].
move: (p ig)=> bh.
have sf: (forall b, image_by_fun g b = inv_image_by_fun (inverse_fun h) b).
move=> a.
rewrite /inv_image_by_fun /inv_image_by_graph /inverse_fun; aw.
by rewrite inverse_graph_involutive /h /restriction2; aw; try ue.
rewrite sf inv_image_of_intersection; last by apply (bijective_inv_function bh).
apply intersectiont_exten => i1; symmetry; apply sf.
Qed.
Theorem complementary_union: forall (I:Set) (f:I-> Set) x,
(forall i, sub (f i) x) -> nonempty I ->
complement x (uniont f) = intersectiont (fun i=> complement x (f i)).
Proof.
move=>I f x als neI; set_extens z; srw.
by move=> [xz nez] j; srw; split=>//; dneg t; exists j.
move: neI=> [_ [i _]] ajz; move: (ajz i); srw; move=> [zx nzf]; split=>//.
by case=> y py; move: (ajz y); srw; intuition.
Qed.
Theorem complementary_intersection: forall (I:Set) (f:I-> Set) x,
(forall i, sub (f i) x) -> nonempty I ->
complement x (intersectiont f) = uniont (fun i=> complement x (f i)).
Proof.
move=> I f x afx neI.
set (g:= fun i => complement x (f i)).
have agw: (forall i, sub (g i) x) by rewrite/g=> i; apply sub_complement.
move:(complementary_union agw neI).
have p: (forall i, complement x (g i) = f i) by rewrite/g=>i; srw.
rewrite- (intersectiont_exten p).
by move => <-; srw; apply sub_uniont.
Qed.
Lemma complementary_union1: forall sf f x,
nonempty sf -> (forall i, inc i sf -> sub (f i) x) ->
complement x (unionf sf f) = intersectionf sf (fun i=> complement x (f i)).
Proof.
move=> sf f x nesf aif; rewrite /unionf/intersectionf -complementary_union//.
by move=> i; apply (aif (Ro i)); apply R_inc.
Qed.
Lemma complementary_intersection1: forall sf f x,
nonempty sf -> (forall i, inc i sf -> sub (f i) x) ->
complement x (intersectionf sf f) = unionf sf (fun i=> complement x (f i)).
Proof.
move=> sf f x nesf hyp.
rewrite /unionf/intersectionf complementary_intersection //.
by move=> i; apply (hyp (Ro i)); apply R_inc.
Qed.
Lemma union_of_twosets_aux: forall x y f,
unionf (doubleton x y) f = union2 (f x) (f y).
Proof.
move=> x y f; set_extens t; aw.
by srw; move=> [z [td tf]]; case (doubleton_or td); move=> <-; intuition.
case=> h; union_tac.
Qed.
Lemma intersection_of_twosets_aux: forall x y f,
intersectionf (doubleton x y) f = intersection2 (f x) (f y).
Proof.
move=> x y f; move: (nonempty_doubleton x y) => ned.
set_extens z; srw; aw; first by move =>h;ee.
by move=> [h1 h2]j; aw;case => ->.
Qed.
Lemma union_of_twosets: forall x y,
unionf (doubleton x y) id = union2 x y.
Proof. by move=> x y;rewrite union_of_twosets_aux. Qed.
Lemma intersection_of_twosets: forall x y,
intersectionf(doubleton x y) id = intersection2 x y.
Proof. by move=> x y; rewrite intersection_of_twosets_aux. Qed.
Lemma intersection_union_distrib1: forall x y z,
union2 x (intersection2 y z) = intersection2 (union2 x y) (union2 x z).
Proof. move=> x y z; set_extens t; aw; intuition. Qed.
Lemma intersection_union_distrib2: forall x y z,
intersection2 x (union2 y z) = union2 (intersection2 x y)(intersection2 x z).
Proof. move=> x y z; set_extens t; aw; intuition. Qed.
Lemma union2_comp: forall x y z,
complement z (union2 x y) = intersection2 (complement z x)(complement z y).
Proof.
move=> x y z; set_extens t; aw; srw; aw; intuition.
Qed.
Lemma intersection2_comp: forall x y z,
sub x z -> sub y z ->
complement z (intersection2 x y) = union2 (complement z x)(complement z y).
Proof.
move=> x y z xz yz; set_extens t; aw; srw; aw; last by intuition.
move=> [te ne].
case (inc_or_not t x) => h; last by auto.
by right; split =>// ; dneg w.
Qed.
Lemma union2_complement: forall x z,
sub x z -> union2 x (complement z x) = z.
Proof.
move=> x z xz; set_extens t; aw.
by case; first (by apply xz); srw; intuition.
move=>tz; case (inc_or_not t (complement z x)); first by intuition.
by move=> h; left; apply use_complement with z.
Qed.
Lemma intersection2_complement: forall x z,
sub x z -> intersection2 x (complement z x) = emptyset.
Proof. by move=> x z xz; empty_tac y ye; move: ye; aw; srw; intuition.
Qed.
Lemma image_of_union2: forall g x y,
is_correspondence g ->
image_by_fun g (union2 x y) = union2 (image_by_fun g x)(image_by_fun g y).
Proof.
move=> g x y cg;
by rewrite -union_of_twosets_aux -union_of_twosets image_of_union.
Qed.
Lemma image_of_intersection2: forall g x y,
is_correspondence g ->
sub (image_by_fun g (intersection2 x y))
(intersection2 (image_by_fun g x) (image_by_fun g y)).
Proof.
move=> g x y cg.
rewrite -intersection_of_twosets -(intersection_of_twosets_aux x y _).
by apply (image_of_intersection).
Qed.
Lemma inv_image_of_intersection2: forall g x y,
is_function g ->
inv_image_by_fun g (intersection2 x y) =
intersection2 (inv_image_by_fun g x)(inv_image_by_fun g y).
Proof.
move=> g x y fg; rewrite - intersection_of_twosets;
rewrite - (intersection_of_twosets_aux x y _).
by apply (inv_image_of_intersection).
Qed.
Lemma inj_image_of_intersection2: forall g x y,
injection g ->
image_by_fun g (intersection2 x y)
= intersection2 (image_by_fun g x)(image_by_fun g y).
Proof.
move=> g x y ig; rewrite - intersection_of_twosets;
rewrite - (intersection_of_twosets_aux x y _).
by apply (inj_image_of_intersection).
Qed.
Lemma inv_image_of_comp: forall f x,
is_function f -> sub x (target f) ->
inv_image_by_fun f (complement (target f) x) =
complement (inv_image_by_fun f (target f))(inv_image_by_fun f x).
Proof.
move=> f x ff sxtf.
rewrite /inv_image_by_fun; set_extens t.
aw;move=> [u []]; srw; aw; move=> [ut nux] Jg.
split; first by ex_tac.
by dneg aa; move: aa=>[v [vx Jvg]]; rewrite -(W_pr ff Jg) (W_pr ff Jvg).
srw;aw; move=> [[ u [ut Jg]] ne]; exists u; split =>//.
by srw; split =>//; dneg h; ex_tac.
Qed.
Lemma inj_image_of_comp: forall f x,
injection f -> sub x (source f) ->
image_by_fun f (complement (source f) x) =
complement (image_by_fun f (source f))(image_by_fun f x).
Proof.
move=> f x inf xsf.
rewrite /image_by_fun; set_extens t; srw; aw.
move=> [z]; srw; aw; move => [[zf nzx] Jg]; split; first by ex_tac.
dneg h; move: h; move=> [z' [z'x] Jg'].
by move: (injective_pr inf Jg Jg') => ->.
move=> [[z [zs Jg]] ne]; ex_tac; srw; split=>//.
by dneg h; ex_tac.
Qed.
We define a covering for a map, a function, and a set of sets
Definition covering f x := fgraph f & sub x (unionb f).
Definition covering_f sf f x := sub x (unionf sf f).
Definition covering_s f x := sub x (union f).
Lemma covering_pr: forall f x,
fgraph f -> covering f x = covering_s (range f) x.
Proof.
move => f x fgf; rewrite /covering/covering_s unionb_alt //.
by apply iff_eq; intuition.
Qed.
Lemma covering_f_pr: forall sf f x,
covering_f sf f x = covering_s (range (L sf f)) x.
Proof.
move=> sf f x.
rewrite /covering_f /covering_s -unionb_alt /unionb; last by gprops.
bw; rewrite (unionf_exten (f':= (fun a : Set => V a (L sf f)))) //.
by move=> i isf; bw.
Qed.
Two definitions, second one will be used in set5
Definition coarser_covering sf f sg g :=
forall j, inc j sg -> exists i, inc i sf & sub (g j) (f i).
Definition coarser_c y y':=
forall a, inc a y' -> exists b, inc b y & sub a b.
Lemma coarser_same: forall sf f sg g ,
coarser_covering sf f sg g = coarser_c (range (L sf f))(range (L sg g)).
Proof.
move=> sf f sg g; rewrite /coarser_c /coarser_covering.
apply iff_eq; move=> cc j.
bw; move=> [b [bsg]] <-; move: (cc _ bsg)=> [i [isf sgjfi]].
have eq: (inc (f i) (range (L sf f))) by bw; ex_tac.
by ex_tac.
move=> jsh.
have gir: (inc (g j) (range (L sg g))) by bw; ex_tac.
move: (cc _ gir)=> [b [br sab]].
move: br; bw; move=> [b' [b'sf eq]]; ex_tac; ue.
Qed.
Lemma sub_covering: forall Ia Ib f x,
covering_f Ia f x -> covering_f Ib f x -> sub Ib Ia ->
coarser_covering Ia f Ib f.
Proof.
by move => Ia Ib f x c1 c2 sba i ib; exists i; split; [ apply sba| fprops].
Qed.
Definition intersection_covering f g :=
fun z => intersection2 (f (P z)) (g (Q z)).
Definition intersection_covering2 x y:=
range (L (product x y)(intersection_covering id id)).
Lemma intersection_covering2_pr: forall x y z,
inc z (intersection_covering2 x y) =
exists a, exists b, inc a x & inc b y & intersection2 a b = z.
Proof.
move=> x y z; rewrite /intersection_covering2 /intersection_covering.
apply iff_eq.
bw; move=> [b [bp iz]]; awi bp.
by exists (P b); exists (Q b); intuition.
by move=> [a [b [ xa [xb iz]]]]; bw; exists (J a b); split; fprops; aw.
Qed.
Lemma intersection_is_covering: forall sf f sg g x,
covering_f sf f x -> covering_f sg g x ->
covering_f (product sf sg) (intersection_covering f g) x.
Proof.
move=> sf f sg g x c1 c2 u ux.
move: (c1 _ ux) (c2 _ ux). srw.
move=> [a [af uf]] [b [bg ug]].
exists (J a b); split; fprops.
by rewrite /intersection_covering; aw.
Qed.
Lemma intersection_covering_coarser1: forall sf f sg g x,
covering_f sf f x -> covering_f sg g x ->
coarser_covering sf f (product sf sg) (intersection_covering f g).
Proof.
move => sf f sg g x c1 c2 u; aw; move=> [pu [pf qg]].
by ex_tac; rewrite /intersection_covering; move => v; aw; intuition.
Qed.
Lemma intersection_covering_coarser2: forall sf f sg g x,
covering_f sf f x -> covering_f sg g x ->
coarser_covering sg g (product sf sg) (intersection_covering f g).
Proof.
move => sf f sg g x c1 c2 u; aw; move=> [pu [pf qg]].
by ex_tac; rewrite /intersection_covering; move => v; aw; intuition.
Qed.
Lemma intersection_covering_coarser3: forall sf f sg g sh h x,
covering_f sf f x -> covering_f sg g x -> covering_f sh h x ->
coarser_covering sf f sh h ->
coarser_covering sg g sh h ->
coarser_covering (product sf sg) (intersection_covering f g) sh h.
Proof.
move => sf f sg g sh h x c1 c2 c3 cc1 cc2 u ush.
move: (cc1 _ ush); move: (cc2 _ ush); move=> [i [isg s1]] [j [jsf s2]].
exists (J j i); split; fprops.
rewrite /intersection_covering; move=> v vh; aw.
by split; [apply s2 |apply s1].
Qed.
Next four lemmas are a rewrite of above; they say that
intersection_covering2 is the supremum of both args
Lemma product_is_covering2: forall u v x,
covering_s u x -> covering_s v x ->
covering_s (intersection_covering2 u v) x.
Proof.
move=> u v x c1 c2 z zu.
move: (c1 _ zu); srw; move=>[y [zy yu]].
move: (c2 _ zu); srw; move=>[y' [zy' yv]].
exists (intersection2 y y'); split; first by aw.
rewrite intersection_covering2_pr.
exists y; exists y'; intuition.
Qed.
Lemma intersection_cov_coarser1: forall f g x,
covering_s f x -> covering_s g x ->
coarser_c f (intersection_covering2 f g).
Proof.
move=> f g x c1 c2 z; rewrite intersection_covering2_pr.
move=> [a [b [af [bg]]]] <-; ex_tac; apply intersection2sub_first.
Qed.
Lemma intersection_cov_coarser2: forall f g x,
covering_s f x -> covering_s g x ->
coarser_c g (intersection_covering2 f g).
Proof.
move=> f g x c1 c2 z; rewrite intersection_covering2_pr.
move=> [a [b [af [bg]]]] <-; ex_tac; apply intersection2sub_second.
Qed.
Lemma intersection_cov_coarser3: forall f g h x,
covering_s f x -> covering_s g x -> covering_s h x ->
coarser_c f h ->
coarser_c g h ->
coarser_c (intersection_covering2 f g) h.
Proof.
move=> f g h x c1 c2 c3 cc1 cc2 t th.
move: (cc1 _ th)=> [b [bf tb]]; move: (cc2 _ th)=> [b' [bg tb']].
exists (intersection2 b b'); split.
by rewrite intersection_covering2_pr; exists b; exists b'; intuition.
move=> z zt; aw; auto.
Qed.
Lemma image_of_covering: forall sf f g,
surjection g -> covering_f sf f (source g)
-> covering_f sf (fun w => image_by_fun g (f w)) (target g).
Proof.
move=> sf f g sg c x xtg.
move: (surjective_pr sg xtg)=> [y [ys Vg]].
move:(unionf_exists (c _ ys)) => [z [zsf yfz]].
union_tac; rewrite /image_by_fun; aw; ex_tac.
Qed.
Lemma inv_image_of_covering: forall sf f g,
is_function g -> covering_f sf f (target g)
-> covering_f sf (fun w => inv_image_by_fun g (f w)) (source g).
Proof.
move=> sf f g fg c x xs.
have Wt: (inc (W x g) (target g)) by fprops.
move: (c _ Wt); srw; move=> [y [ysf Wf]].
rewrite /inv_image_by_fun; ex_tac; aw; ex_tac; graph_tac.
Qed.
Lemma product_of_covering: forall sf f sg g x y,
covering_f sf f x -> covering_f sg g y ->
covering_f (product sf sg) (fun z => product (f (P z)) (g (Q z)))
(product x y).
Proof.
move=> sf f sg g x y c1 c2 z; aw; move=> [zp [px qy]].
move: (c1 _ px) (c2 _ qy); srw.
by move=> [a [asf pa]][b [bsg qb]]; exists (J a b); aw; intuition.
Qed.
Lemma agrees_on_covering: forall sf f x g g',
covering_f sf f x -> is_function g -> is_function g' ->
source g = x -> source g' = x ->
(forall i, inc i sf -> agrees_on (intersection2 x (f i)) g g') ->
agrees_on x g g'.
Proof.
move=> sf f x g g' c fg fg' sg sg' ag.
hnf; rewrite sg sg';split; fprops; split; fprops.
move=> a ax; move: (c _ ax); srw; move => [y [ysf afy]].
by move: (ag _ ysf)=> [ _ [ _ ]]; apply; apply intersection2_inc.
Qed.
Definition function_prop f s t:=
is_function f & source f = s & target f = t.
Definition function_prop_sub f s t:=
is_function f & source f = s & sub (target f) t.
Lemma extension_covering1: forall sf f t h,
(forall i, inc i sf -> function_prop (h i)(f i) t ) ->
(forall i j, inc i sf -> inc j sf ->
agrees_on (intersection2 (f i) (f j)) (h i) (h j)) ->
exists g, function_prop g (unionf sf f) t &
graph g = (unionf sf (fun i => (graph (h i)))) &
range(graph g) = (unionf sf (fun i => (range (graph (h i))))) &
(forall i, inc i sf -> agrees_on (f i) g (h i)).
Proof.
move=> sf f t h afp aag.
set (g:= corresp (unionf sf f) t (unionf sf (fun i => (graph (h i))))).
have sg: (source g = unionf sf f) by rewrite /g; aw.
have tg: (target g = t) by rewrite /g; aw.
have gg: (graph g = (unionf sf (fun i => graph (h i)))) by rewrite /g; aw.
have gu: (is_graph (unionf sf (fun i => graph (h i)))).
move=> z; srw; move=> [y [ysf zg]].
by move: (afp _ ysf) => [ff _]; move: (function_graph ff); apply.
have fgf: fgraph (unionf sf (fun i : Set => graph (h i))).
split=>//.
move => x y; srw; move=> [u [usf xghu]] [v [vsf yghv]] sp.
move:(afp _ usf) => [fu [shu _]]; move:(afp _ vsf) => [fv [shv _]].
move: (W_pr2 fu xghu) => qx; move: (W_pr2 fv yghv)=>qy.
apply pair_exten=>//.
by apply (function_graph fu).
by apply (function_graph fv).
rewrite qx qy sp; move: (aag _ _ usf vsf)=> [_[_]]; apply.
by rewrite -shu -shv; apply intersection2_inc ;[ rewrite -sp|]; graph_tac.
have rgg: range (graph g) = unionf sf (fun i => range (graph (h i))).
rewrite /g; aw; set_extens zs.
dw; move=> [x]; srw; move=> [y [ysf zfy]].
move: (afp _ ysf) => [fhy [ shy _]]; ex_tac; dw; fprops.
by exists x.
rewrite unionf_rw; move=> [y [ysf]].
move: (afp _ ysf) => [fhy [ shy _]]; dw; fprops.
by move=> [x Jg]; exists x; srw; ex_tac; fprops.
have fg: is_function g.
rewrite /g; apply is_function_pr=>//.
move=> z; dw; move=> [x]; srw; move=> [y [ysf Jg]].
by move: (afp _ ysf)=> [fhy [_]] <-; graph_tac.
set_extens z; dw.
srw; move=> [y [ysf zfy]].
move: (afp _ ysf) => [fhy [ shy _]].
by exists (W z (h y)); union_tac; graph_tac; ue.
move=> [y]; srw; move=> [x [xsf Jg]].
move: (afp _ xsf) => [fhx [shx _]].
by ex_tac; rewrite -shx; graph_tac.
exists g; split; first (by hnf); split =>//; split =>//.
move=> y ysf; move: (afp _ ysf) => [fhy [ shy _]].
split.
by rewrite /g; aw; move=> x xfy; union_tac.
split; first by rewrite shy; fprops.
move=> a afy.
have aux: (inc (J a (W a (h y))) (graph g)).
by rewrite gg; union_tac; graph_tac; rewrite shy.
by apply W_pr; fprops.
Qed.
Lemma extension_covering: forall sf f t h,
(forall i, inc i sf -> function_prop (h i) (f i) t) ->
(forall i j, inc i sf -> inc j sf ->
agrees_on (intersection2 (f i) (f j)) (h i) (h j)) ->
exists_unique (fun g => function_prop g (unionf sf f) t &
(forall i, inc i sf -> agrees_on (f i) g (h i))).
Proof.
move=> sf f t h afp aag; split.
move: (extension_covering1 afp aag)=> [g [p1 [p2 [p3 p4]]]].
by exists g; intuition.
rewrite /function_prop=> x y [[fx [sx tx]] agx] [[fy [sy ty]] agy].
apply function_exten=>//; try ue.
have c: (covering_f sf f (unionf sf f)) by hnf; fprops.
have agi: (forall i,
inc i sf -> agrees_on (intersection2 (unionf sf f) (f i)) x y).
move=> i isf; hnf.
split;first by rewrite sx;apply (intersection2sub_first).
split;first by rewrite sy;apply (intersection2sub_first).
move => a; aw; move: (agx _ isf) (agy _ isf).
rewrite /agrees_on; move=> [_ [_ p1]][_ [_ p2]] [_ afi].
by rewrite (p1 _ afi) (p2 _ afi).
move:(agrees_on_covering c fx fy sx sy agi)=>p.
move:p => [_[_]]; rewrite sx; apply.
Qed.
covering_s u x -> covering_s v x ->
covering_s (intersection_covering2 u v) x.
Proof.
move=> u v x c1 c2 z zu.
move: (c1 _ zu); srw; move=>[y [zy yu]].
move: (c2 _ zu); srw; move=>[y' [zy' yv]].
exists (intersection2 y y'); split; first by aw.
rewrite intersection_covering2_pr.
exists y; exists y'; intuition.
Qed.
Lemma intersection_cov_coarser1: forall f g x,
covering_s f x -> covering_s g x ->
coarser_c f (intersection_covering2 f g).
Proof.
move=> f g x c1 c2 z; rewrite intersection_covering2_pr.
move=> [a [b [af [bg]]]] <-; ex_tac; apply intersection2sub_first.
Qed.
Lemma intersection_cov_coarser2: forall f g x,
covering_s f x -> covering_s g x ->
coarser_c g (intersection_covering2 f g).
Proof.
move=> f g x c1 c2 z; rewrite intersection_covering2_pr.
move=> [a [b [af [bg]]]] <-; ex_tac; apply intersection2sub_second.
Qed.
Lemma intersection_cov_coarser3: forall f g h x,
covering_s f x -> covering_s g x -> covering_s h x ->
coarser_c f h ->
coarser_c g h ->
coarser_c (intersection_covering2 f g) h.
Proof.
move=> f g h x c1 c2 c3 cc1 cc2 t th.
move: (cc1 _ th)=> [b [bf tb]]; move: (cc2 _ th)=> [b' [bg tb']].
exists (intersection2 b b'); split.
by rewrite intersection_covering2_pr; exists b; exists b'; intuition.
move=> z zt; aw; auto.
Qed.
Lemma image_of_covering: forall sf f g,
surjection g -> covering_f sf f (source g)
-> covering_f sf (fun w => image_by_fun g (f w)) (target g).
Proof.
move=> sf f g sg c x xtg.
move: (surjective_pr sg xtg)=> [y [ys Vg]].
move:(unionf_exists (c _ ys)) => [z [zsf yfz]].
union_tac; rewrite /image_by_fun; aw; ex_tac.
Qed.
Lemma inv_image_of_covering: forall sf f g,
is_function g -> covering_f sf f (target g)
-> covering_f sf (fun w => inv_image_by_fun g (f w)) (source g).
Proof.
move=> sf f g fg c x xs.
have Wt: (inc (W x g) (target g)) by fprops.
move: (c _ Wt); srw; move=> [y [ysf Wf]].
rewrite /inv_image_by_fun; ex_tac; aw; ex_tac; graph_tac.
Qed.
Lemma product_of_covering: forall sf f sg g x y,
covering_f sf f x -> covering_f sg g y ->
covering_f (product sf sg) (fun z => product (f (P z)) (g (Q z)))
(product x y).
Proof.
move=> sf f sg g x y c1 c2 z; aw; move=> [zp [px qy]].
move: (c1 _ px) (c2 _ qy); srw.
by move=> [a [asf pa]][b [bsg qb]]; exists (J a b); aw; intuition.
Qed.
Lemma agrees_on_covering: forall sf f x g g',
covering_f sf f x -> is_function g -> is_function g' ->
source g = x -> source g' = x ->
(forall i, inc i sf -> agrees_on (intersection2 x (f i)) g g') ->
agrees_on x g g'.
Proof.
move=> sf f x g g' c fg fg' sg sg' ag.
hnf; rewrite sg sg';split; fprops; split; fprops.
move=> a ax; move: (c _ ax); srw; move => [y [ysf afy]].
by move: (ag _ ysf)=> [ _ [ _ ]]; apply; apply intersection2_inc.
Qed.
Definition function_prop f s t:=
is_function f & source f = s & target f = t.
Definition function_prop_sub f s t:=
is_function f & source f = s & sub (target f) t.
Lemma extension_covering1: forall sf f t h,
(forall i, inc i sf -> function_prop (h i)(f i) t ) ->
(forall i j, inc i sf -> inc j sf ->
agrees_on (intersection2 (f i) (f j)) (h i) (h j)) ->
exists g, function_prop g (unionf sf f) t &
graph g = (unionf sf (fun i => (graph (h i)))) &
range(graph g) = (unionf sf (fun i => (range (graph (h i))))) &
(forall i, inc i sf -> agrees_on (f i) g (h i)).
Proof.
move=> sf f t h afp aag.
set (g:= corresp (unionf sf f) t (unionf sf (fun i => (graph (h i))))).
have sg: (source g = unionf sf f) by rewrite /g; aw.
have tg: (target g = t) by rewrite /g; aw.
have gg: (graph g = (unionf sf (fun i => graph (h i)))) by rewrite /g; aw.
have gu: (is_graph (unionf sf (fun i => graph (h i)))).
move=> z; srw; move=> [y [ysf zg]].
by move: (afp _ ysf) => [ff _]; move: (function_graph ff); apply.
have fgf: fgraph (unionf sf (fun i : Set => graph (h i))).
split=>//.
move => x y; srw; move=> [u [usf xghu]] [v [vsf yghv]] sp.
move:(afp _ usf) => [fu [shu _]]; move:(afp _ vsf) => [fv [shv _]].
move: (W_pr2 fu xghu) => qx; move: (W_pr2 fv yghv)=>qy.
apply pair_exten=>//.
by apply (function_graph fu).
by apply (function_graph fv).
rewrite qx qy sp; move: (aag _ _ usf vsf)=> [_[_]]; apply.
by rewrite -shu -shv; apply intersection2_inc ;[ rewrite -sp|]; graph_tac.
have rgg: range (graph g) = unionf sf (fun i => range (graph (h i))).
rewrite /g; aw; set_extens zs.
dw; move=> [x]; srw; move=> [y [ysf zfy]].
move: (afp _ ysf) => [fhy [ shy _]]; ex_tac; dw; fprops.
by exists x.
rewrite unionf_rw; move=> [y [ysf]].
move: (afp _ ysf) => [fhy [ shy _]]; dw; fprops.
by move=> [x Jg]; exists x; srw; ex_tac; fprops.
have fg: is_function g.
rewrite /g; apply is_function_pr=>//.
move=> z; dw; move=> [x]; srw; move=> [y [ysf Jg]].
by move: (afp _ ysf)=> [fhy [_]] <-; graph_tac.
set_extens z; dw.
srw; move=> [y [ysf zfy]].
move: (afp _ ysf) => [fhy [ shy _]].
by exists (W z (h y)); union_tac; graph_tac; ue.
move=> [y]; srw; move=> [x [xsf Jg]].
move: (afp _ xsf) => [fhx [shx _]].
by ex_tac; rewrite -shx; graph_tac.
exists g; split; first (by hnf); split =>//; split =>//.
move=> y ysf; move: (afp _ ysf) => [fhy [ shy _]].
split.
by rewrite /g; aw; move=> x xfy; union_tac.
split; first by rewrite shy; fprops.
move=> a afy.
have aux: (inc (J a (W a (h y))) (graph g)).
by rewrite gg; union_tac; graph_tac; rewrite shy.
by apply W_pr; fprops.
Qed.
Lemma extension_covering: forall sf f t h,
(forall i, inc i sf -> function_prop (h i) (f i) t) ->
(forall i j, inc i sf -> inc j sf ->
agrees_on (intersection2 (f i) (f j)) (h i) (h j)) ->
exists_unique (fun g => function_prop g (unionf sf f) t &
(forall i, inc i sf -> agrees_on (f i) g (h i))).
Proof.
move=> sf f t h afp aag; split.
move: (extension_covering1 afp aag)=> [g [p1 [p2 [p3 p4]]]].
by exists g; intuition.
rewrite /function_prop=> x y [[fx [sx tx]] agx] [[fy [sy ty]] agy].
apply function_exten=>//; try ue.
have c: (covering_f sf f (unionf sf f)) by hnf; fprops.
have agi: (forall i,
inc i sf -> agrees_on (intersection2 (unionf sf f) (f i)) x y).
move=> i isf; hnf.
split;first by rewrite sx;apply (intersection2sub_first).
split;first by rewrite sy;apply (intersection2sub_first).
move => a; aw; move: (agx _ isf) (agy _ isf).
rewrite /agrees_on; move=> [_ [_ p1]][_ [_ p2]] [_ afi].
by rewrite (p1 _ afi) (p2 _ afi).
move:(agrees_on_covering c fx fy sx sy agi)=>p.
move:p => [_[_]]; rewrite sx; apply.
Qed.
We have two definitions for a partition; in one case, we add the constraint
that elements are nonempty
Definition mutually_disjoint f :=
(forall i j, inc i (domain f) -> inc j (domain f) ->
i = j \/ (disjoint (V i f) (V j f))).
Definition partition_s y x:=
(union y = x) &
(forall a b, inc a y -> inc b y -> a = b \/ disjoint a b).
Definition partition y x:=
(union y = x) &
(forall a, inc a y -> nonempty a) &
(forall a b, inc a y -> inc b y -> a = b \/ disjoint a b).
Definition partition_fam f x:=
fgraph f & mutually_disjoint f & unionb f = x.
Ltac eq_dichot v:= match goal with |- ?a = ?b \/ _
=> case (equal_or_not a b); first (by intuition); move=> v; right end.
Ltac mdi_tac v:= eq_dichot v; apply disjoint_pr.
Lemma mutually_disjoint_prop: forall f,
(forall i j y, inc i (domain f) -> inc j (domain f) ->
inc y (V i f) -> inc y (V j f) -> i=j) ->
mutually_disjoint f.
Proof.
move=> f hyp i j idf jdf; mdi_tac nij => u ui uj; elim nij; exact: (hyp _ _ u).
Qed.
Lemma mutually_disjoint_prop1: forall f, is_function f ->
(forall i j y, inc i (source f) -> inc j (source f) ->
inc y (W i f) -> inc y (W j f) -> i=j) ->
mutually_disjoint (graph f).
Proof.
by move=> f ff hyp; apply mutually_disjoint_prop; rewrite f_domain_graph.
Qed.
Lemma mutually_disjoint_prop2: forall x f,
(forall i j y, inc i x -> inc j x ->
inc y (f i) -> inc y (f j) -> i=j) ->
mutually_disjoint (L x f).
Proof.
move=> x f hyp; apply mutually_disjoint_prop. bw.
by move=> i j y ix jy; bw; apply hyp.
Qed.
Lemma partitionset_pr: forall y x,
partition y x -> partition_s y x.
Proof. rewrite /partition/partition_s; intuition. Qed.
Lemma partition_same: forall y x,
partition_s y x -> partition_fam (identity_g y) x.
Proof.
move=> y x [un di]; hnf; rewrite unionb_identity.
split; first (by apply identity_fgraph); split=>//.
hnf; rewrite identity_domain; move=> i j iy jy.
by do 2 rewrite identity_ev //; apply di.
Qed.
Lemma partition_same2: forall y x,
partition_fam y x -> partition_s (range y) x.
Proof.
move=>y x [fg [md uyx]]; split; first by rewrite -unionb_alt.
move=> a b; rewrite (frange_inc_rw a fg) (frange_inc_rw b fg) .
move=> [u [ud]] -> [v [vd]] ->.
by case (md _ _ ud vd) ;[ move=>->|]; intuition.
Qed.
Lemma partitions_is_covering: forall y x,
partition_s y x -> covering_s y x.
Proof. by move=> y x [u d] a ax; ue. Qed.
Lemma partition_fam_is_covering: forall y x,
partition_fam y x -> covering y x.
Proof. move=> y x [fg [md uyx]];split=>//; rewrite uyx; fprops. Qed.
Lemma partition_inc_exists: forall f x y,
partition_fam f x -> inc y x -> exists i, (inc i (domain f) & inc y (V i f)).
Proof. by move=> f x y [fg [md ufx]] ;rewrite -ufx; srw. Qed.
Lemma partition_inc_unique: forall f x i j y,
partition_fam f x -> inc i (domain f) -> inc y (V i f) ->
inc j (domain f) -> inc y (V j f) -> i = j.
Proof.
move=> f x i j y [fg [md ufx]] id yif jd yjf.
case (md i j id jd)=>//.
have: (inc y (intersection2 (V i f) (V j f))) by apply intersection2_inc.
by rewrite /disjoint; move=> yi di; move: yi; rewrite di; case; case.
Qed.
Lemma coarser_transitive : forall y y' y'',
coarser_c y y' -> coarser_c y' y'' -> coarser_c y y''.
Proof.
move=> y y' y'' c1 c2 u uy.
move: (c2 _ uy)=> [z [zy uz]]; move: (c1 _ zy)=> [t [ty zt]].
by ex_tac; apply (sub_trans uz zt).
Qed.
Lemma coarser_reflexive : forall y, coarser_c y y.
Proof. by move=> y x xy; ex_tac. Qed.
Lemma coarser_antisymmetric : forall y y' x,
partition y x -> partition y' x ->
coarser_c y y' -> coarser_c y' y -> y = y'.
Proof.
suff: (forall y y' x, partition y x -> partition y' x ->
coarser_c y y' -> coarser_c y' y -> sub y y').
move=> hyp y y' x p1 p2 c1 c2.
apply extensionality; first apply (hyp _ _ _ p1 p2 c1 c2).
by apply (hyp _ _ _ p2 p1 c2 c1).
move=> y y' x [uyx [ne adi]] p2 c1 c2 t ty.
move: (c2 _ ty) => [b [biy tb]]; move: (c1 _ biy) => [a [ay ba]].
have sta: sub t a by apply (sub_trans tb ba).
move: (ne _ ty)=> [z zt].
have nd: inc z (intersection2 t a) by apply intersection2_inc=>//; apply sta.
case (adi _ _ ty ay) => eq.
have: t = b by apply extensionality=>//; rewrite eq.
by move =>->.
by red in eq; empty_tac1 z.
Qed.
Definition variant a x y := (fun z:Set => Yo (z = a) x y).
Definition variantL a b x y := L (doubleton a b) (variant a x y).
Lemma variant_if_rw: forall a x y z,
z = a -> variant a x y z = x.
Proof. by move=> a x yz z za; rewrite /variant Y_if_rw. Qed.
Lemma variant_if_not_rw: forall a x y z,
z <> a -> variant a x y z = y.
Proof. by move=> a x y z za; rewrite /variant Y_if_not_rw. Qed.
Definition varianti x a b := fun z => Yo (inc z x) a b.
Lemma varianti_in: forall z x a b, inc z x -> (varianti x a b z) = a.
Proof. by move=> z x a b zx; rewrite /varianti; Ytac0. Qed.
Lemma varianti_out: forall z x a b, ~ inc z x -> (varianti x a b z) = b.
Proof. by move=> z x a b zx; rewrite /varianti; Ytac0. Qed.
Lemma variant_V_a : forall a b x y,
V a (variantL a b x y) = x.
Proof. by rewrite /variantL=> a b x y; bw; fprops; apply variant_if_rw. Qed.
Lemma variant_V_b : forall a b x y,
b <> a -> V b (variantL a b x y) = y.
Proof. by rewrite /variantL=> a b x y; bw; fprops; apply variant_if_not_rw.
Qed.
Lemma variant_fgraph : forall a b x y,
fgraph (variantL a b x y).
Proof. rewrite/variantL=> a b x y; gprops. Qed.
Lemma variant_domain : forall a b x y,
domain (variantL a b x y) = doubleton a b.
Proof. rewrite/variantL=> a b x y; bw. Qed.
Definition variantLc f g:= variantL TPa TPb f g.
Lemma variantLc_fgraph : forall x y,
fgraph (variantLc x y).
Proof. rewrite/variantLc/variantL=> x y; gprops. Qed.
Hint Resolve variant_fgraph variantLc_fgraph: fprops.
Lemma variantLc_domain: forall f g,
domain (variantLc f g) = two_points.
Proof. rewrite/variantLc/variantL=> x y; bw; apply two_points_pr2. Qed.
Hint Rewrite variantLc_domain: bw.
Lemma variantLc_domain_nonempty: forall f g,
nonempty (domain (variantLc f g)).
Proof. move=> f g; bw; exists TPa; fprops. Qed.
Lemma variant_V_ca : forall x y,
V TPa (variantLc x y) = x.
Proof.
by rewrite /variantLc/variantL=> x y; bw;[ apply variant_if_rw |fprops].
Qed.
Lemma variant_V_cb : forall x y,
V TPb (variantLc x y) = y.
Proof.
rewrite /variantLc/variantL=> x y; bw;[ apply variant_if_not_rw |fprops].
apply two_points_distinctb.
Qed.
Hint Rewrite variant_V_ca variant_V_cb: bw.
Ltac try_lvariant u:=
move:u;rewrite {1} two_points_rw; case => ->; bw.
Definition partition_with_complement x j :=
variantLc j (complement x j).
Lemma is_partition_with_complement: forall x j,
sub j x -> partition_fam (partition_with_complement x j) x.
Proof.
rewrite /partition_with_complement=> x j jx; split.
by apply variantLc_fgraph.
split.
move=> a b; bw; aw; case => ->; case => ->; bw; intuition;
right;try apply disjoint_complement.
by apply disjoint_symmetric; apply disjoint_complement.
rewrite/unionb /variantLc variant_domain union_of_twosets_aux.
by bw; apply union2_complement.
Qed.
Greatest and least partition for covering order
Definition largest_partition x := fun_image x singleton.
Definition smallest_partition x := (singleton x).
Lemma partition_smallest: forall x,
nonempty x -> partition(smallest_partition x) x.
Proof.
rewrite /smallest_partition=> x ne.
split; first by apply union_singleton.
split.
by move=> a; aw; move=> ->.
by move=> a b; aw; move=> -> ->; intuition.
Qed.
Lemma largest_partition_pr: forall x z,
inc z (largest_partition x) = exists w, inc w x & singleton w = z.
Proof. by rewrite/largest_partition=> x z; aw.
Qed.
Lemma partition_largest: forall x, partition (largest_partition x) x.
Proof.
move=> x; split.
set_extens t; srw.
move=> [y [yt]]; rewrite largest_partition_pr; move=> [w [wx swy]].
by move: yt; rewrite -swy; aw =>->.
move=> tx; exists (singleton t).
by split; first (by fprops); rewrite largest_partition_pr; ex_tac.
split.
move=> a; rewrite largest_partition_pr;
by move=> [w [wx]] <-; apply nonempty_singleton.
move=> a b; rewrite 2! largest_partition_pr.
move=> [w [wx]] <- [w' [wx']] <-.
by mdi_tac ww=> u; aw => uw vw; apply ww; rewrite -uw -vw.
Qed.
Definition injective_graph f:=
fgraph f &
(forall x y, inc x (domain f) -> inc y (domain f) ->
V x f = V y f -> x = y).
Lemma injective_partition: forall f x,
partition_fam f x -> (forall i, inc i (domain f) -> nonempty (V i f))
-> injective_graph f.
Proof.
move=> f x [fgf [md ux]] hyp; split=>//.
move=> a b ad bd sV; move: (md _ _ ad bd) (hyp _ bd)=> [] //.
rewrite /disjoint sV intersection2idem=> ->.
by move=> ?; elim not_nonempty_empty.
Qed.
Lemma partition_fam_partition: forall f x,
partition_fam f x -> (forall i, inc i (domain f) -> nonempty (V i f))
-> partition(range f) x.
Proof.
move=> f x pf hyp; move: (partition_same2 pf)=> [ur di].
move:pf=> [fgf [md _]].
split=>//;split=>//.
have gf: (is_graph f) by fprops.
move=> a; dw; move=> [b bf]; move: (pr2_V fgf bf); aw =>->.
by apply hyp; ex_tac.
Qed.
Lemma inv_image_disjoint : forall g x y,
is_function g -> disjoint x y ->
disjoint (inv_image_by_fun g x)(inv_image_by_fun g y).
Proof.
move=> g x y fg dxy; rewrite /disjoint -inv_image_of_intersection2 //.
rewrite/inv_image_by_fun; empty_tac t te.
move: te; aw; move=> [u [ui _]]; rewrite dxy in ui; empty_tac0.
Qed.
Theorem extension_partition: forall f x t h,
partition_fam f x ->
(forall i, inc i (domain f) -> function_prop (h i) (V i f) t) ->
exists_unique (fun g => function_prop g x t &
(forall i, inc i (domain f) -> agrees_on (V i f) g (h i))).
Proof.
move=> f x t h [fgf [mdf ufx]] afp.
set (ff:= fun i => V i f).
suff aah: (forall i j, inc i (domain f) -> inc j (domain f) ->
agrees_on (intersection2 (ff i) (ff j)) (h i) (h j)).
by move: (extension_covering afp aah); rewrite -ufx.
move=> i j id jd; case(mdf _ _ id jd).
move=>->; rewrite intersection2idem.
move: (afp _ jd)=> [_ [shj _]].
by red; rewrite shj; intuition.
rewrite/disjoint=> ->.
red; split; try split; try apply emptyset_sub_any.
by move=> a ae; elim (emptyset_pr ae).
Qed.
Theorem extension_partition1: forall f x t h,
partition_fam f x ->
(forall i, inc i (domain f) -> function_prop_sub (h i) (V i f) t) ->
exists g, function_prop g x t &
(forall i, inc i (domain f) -> agrees_on (V i f) g (h i)).
Proof.
move=> f x t h pf prop.
pose nh i := corresp (V i f) t (graph (h i)).
have afp:(forall i, inc i (domain f) -> function_prop (nh i) (V i f) t).
move=> i id; rewrite /nh /function_prop; aw; split ;last by [].
move: (prop _ id); move=> [fgh [sg th]].
apply is_function_pr; fprops.
by apply sub_trans with (target (h i))=>//; apply f_range_graph =>//.
by rewrite f_domain_graph.
move: (extension_partition pf afp)=> [[g [fpg aag]] _].
exists g; split=>//.
move=> i idf; move: (aag _ idf)=> [ aa [ cc bb]].
split=>//; split.
by move: (prop _ idf)=> [_ [ shi _]]; rewrite shi; fprops.
by move=> a aV; move: (bb _ aV); rewrite /nh=>->; rewrite /W; aw.
Qed.
Definition smallest_partition x := (singleton x).
Lemma partition_smallest: forall x,
nonempty x -> partition(smallest_partition x) x.
Proof.
rewrite /smallest_partition=> x ne.
split; first by apply union_singleton.
split.
by move=> a; aw; move=> ->.
by move=> a b; aw; move=> -> ->; intuition.
Qed.
Lemma largest_partition_pr: forall x z,
inc z (largest_partition x) = exists w, inc w x & singleton w = z.
Proof. by rewrite/largest_partition=> x z; aw.
Qed.
Lemma partition_largest: forall x, partition (largest_partition x) x.
Proof.
move=> x; split.
set_extens t; srw.
move=> [y [yt]]; rewrite largest_partition_pr; move=> [w [wx swy]].
by move: yt; rewrite -swy; aw =>->.
move=> tx; exists (singleton t).
by split; first (by fprops); rewrite largest_partition_pr; ex_tac.
split.
move=> a; rewrite largest_partition_pr;
by move=> [w [wx]] <-; apply nonempty_singleton.
move=> a b; rewrite 2! largest_partition_pr.
move=> [w [wx]] <- [w' [wx']] <-.
by mdi_tac ww=> u; aw => uw vw; apply ww; rewrite -uw -vw.
Qed.
Definition injective_graph f:=
fgraph f &
(forall x y, inc x (domain f) -> inc y (domain f) ->
V x f = V y f -> x = y).
Lemma injective_partition: forall f x,
partition_fam f x -> (forall i, inc i (domain f) -> nonempty (V i f))
-> injective_graph f.
Proof.
move=> f x [fgf [md ux]] hyp; split=>//.
move=> a b ad bd sV; move: (md _ _ ad bd) (hyp _ bd)=> [] //.
rewrite /disjoint sV intersection2idem=> ->.
by move=> ?; elim not_nonempty_empty.
Qed.
Lemma partition_fam_partition: forall f x,
partition_fam f x -> (forall i, inc i (domain f) -> nonempty (V i f))
-> partition(range f) x.
Proof.
move=> f x pf hyp; move: (partition_same2 pf)=> [ur di].
move:pf=> [fgf [md _]].
split=>//;split=>//.
have gf: (is_graph f) by fprops.
move=> a; dw; move=> [b bf]; move: (pr2_V fgf bf); aw =>->.
by apply hyp; ex_tac.
Qed.
Lemma inv_image_disjoint : forall g x y,
is_function g -> disjoint x y ->
disjoint (inv_image_by_fun g x)(inv_image_by_fun g y).
Proof.
move=> g x y fg dxy; rewrite /disjoint -inv_image_of_intersection2 //.
rewrite/inv_image_by_fun; empty_tac t te.
move: te; aw; move=> [u [ui _]]; rewrite dxy in ui; empty_tac0.
Qed.
Theorem extension_partition: forall f x t h,
partition_fam f x ->
(forall i, inc i (domain f) -> function_prop (h i) (V i f) t) ->
exists_unique (fun g => function_prop g x t &
(forall i, inc i (domain f) -> agrees_on (V i f) g (h i))).
Proof.
move=> f x t h [fgf [mdf ufx]] afp.
set (ff:= fun i => V i f).
suff aah: (forall i j, inc i (domain f) -> inc j (domain f) ->
agrees_on (intersection2 (ff i) (ff j)) (h i) (h j)).
by move: (extension_covering afp aah); rewrite -ufx.
move=> i j id jd; case(mdf _ _ id jd).
move=>->; rewrite intersection2idem.
move: (afp _ jd)=> [_ [shj _]].
by red; rewrite shj; intuition.
rewrite/disjoint=> ->.
red; split; try split; try apply emptyset_sub_any.
by move=> a ae; elim (emptyset_pr ae).
Qed.
Theorem extension_partition1: forall f x t h,
partition_fam f x ->
(forall i, inc i (domain f) -> function_prop_sub (h i) (V i f) t) ->
exists g, function_prop g x t &
(forall i, inc i (domain f) -> agrees_on (V i f) g (h i)).
Proof.
move=> f x t h pf prop.
pose nh i := corresp (V i f) t (graph (h i)).
have afp:(forall i, inc i (domain f) -> function_prop (nh i) (V i f) t).
move=> i id; rewrite /nh /function_prop; aw; split ;last by [].
move: (prop _ id); move=> [fgh [sg th]].
apply is_function_pr; fprops.
by apply sub_trans with (target (h i))=>//; apply f_range_graph =>//.
by rewrite f_domain_graph.
move: (extension_partition pf afp)=> [[g [fpg aag]] _].
exists g; split=>//.
move=> i idf; move: (aag _ idf)=> [ aa [ cc bb]].
split=>//; split.
by move: (prop _ idf)=> [_ [ shi _]]; rewrite shi; fprops.
by move=> a aV; move: (bb _ aV); rewrite /nh=>->; rewrite /W; aw.
Qed.
Theorem disjoint_union_lemma : forall f,
fgraph f -> exists g, exists x,
fgraph g & x = unionb g &
(forall i, inc i (domain f) -> equipotent (V i f) (V i g))
& mutually_disjoint g.
Proof.
move=> f fgf.
set (a:= unionb f); set (sf:= domain f).
pose h (i:sf) := BL(fun w =>J w (Ro i))(V (Ro i) f)(product a sf).
have shi: (forall i:sf, source (h i) = V (Ro i) f) by rewrite /h=>i; aw.
have thi: (forall i, target (h i) = product a sf) by rewrite /h=>i; aw.
have ahi:forall i:sf, transf_axioms (fun w => J w (Ro i)) (V (Ro i) f)
(product a sf).
move=> i v vs; apply product_pair_inc; last by apply R_inc.
by apply unionb_inc with (Ro i)=>//; apply R_inc.
have fhi: (forall i:sf, is_function (h i)) by rewrite /h=>i; apply bl_function.
set (g := IM (fun i:sf => J (Ro i) (image_of_fun (h i)))).
have fgg: fgraph g.
rewrite/g; split.
by move=> t; rewrite IM_rw; move=> [z] <-; fprops.
move=> x y; rewrite 2! IM_rw; move=> [z1] <- [z2] <-; aw.
by move => sr; rewrite (R_inj sr).
have mdg: mutually_disjoint g.
have Ht: is_graph g by fprops.
move=> x y; dw; move=> [x' Jxg][y' Jyg].
move: (pr2_V fgg Jxg); move: (pr2_V fgg Jyg); aw=> <- <-.
mdi_tac ne => u ux uy; elim ne; move: ux uy.
move: Jxg; rewrite/g; rewrite IM_rw; move=> [x'' p1].
rewrite -(pr1_def p1) - (pr2_def p1) {p1}.
move: Jyg; rewrite/g;rewrite IM_rw; move=> [y'' p2].
rewrite -(pr1_def p2) - (pr2_def p2) {p2}.
have hf1: (is_function (h x'')) by apply bl_function.
have hf2: (is_function (h y'')) by apply bl_function.
rewrite /image_of_fun /h; aw.
move=> [x''' [xs Jxg]] [y''' [ys Jyg]].
move: (W_pr hf1 Jxg); rewrite -(W_pr hf2 Jyg).
by rewrite /h; aw; fprops;apply pr2_def.
exists g; exists (unionb g).
split=>//; split=>//; split=>//.
move=>i isf; set (j:= Bo isf).
have Jg: inc (J (Ro j) (image_of_fun (h j))) g by rewrite/g IM_rw; exists j.
move: (pr2_V fgg Jg); aw; rewrite (B_eq isf) => ih.
have fhj: is_function (h j) by rewrite /h; apply bl_function; apply ahi.
have rah: (restriction2_axioms (h j) (V i f) (image_of_fun (h j))).
split =>//; split; first by rewrite shi (B_eq isf).
split; first by apply sub_image_target=>//.
rewrite/image_of_fun/image_by_fun.
have ->: (source (h j) = V i f) by rewrite/ h (B_eq isf); aw.
fprops.
exists (restriction2 (h j) (V i f) (image_of_fun (h j))).
split; split.
apply restriction2_injective; last by apply rah.
rewrite /h; apply bl_injective; first by apply ahi.
by move=> u v _ _; apply pr1_def.
apply restriction2_surjective=>//; by rewrite shi (B_eq isf).
rewrite/restriction2; aw.
rewrite/restriction2; aw.
Qed.
Definition disjoint_union_fam f :=
L (domain f)(fun i => product (V i f) (singleton i)).
Definition disjoint_union f :=
unionb (disjoint_union_fam f).
Lemma disjoint_union_disjoint:
forall f, fgraph f ->
mutually_disjoint(disjoint_union_fam f).
Proof.
move=> f fgf i j; rewrite /disjoint_union_fam L_domain=> id jd; bw.
by mdi_tac t => u; aw=> [] [_ [_ eq1]][_ [_]]; rewrite eq1.
Qed.
Lemma partion_union_disjoint: forall f, fgraph f ->
partition_fam (disjoint_union_fam f) (disjoint_union f).
Proof.
move=> f fgf; split.
by rewrite /disjoint_union_fam; gprops.
split=>//; by apply disjoint_union_disjoint.
Qed.
Lemma du_index_pr:forall f x, inc x (disjoint_union f) ->
(inc (Q x) (domain f) & inc (P x) (V (Q x) f) & is_pair x).
Proof.
rewrite /disjoint_union/disjoint_union_fam=> f x; srw; bw.
move=> [y [yd]]; bw; aw; move => [xp [pv]] ->; intuition.
Qed.
Lemma inc_disjoint_union: forall f x y,
inc y (domain f) -> inc x (V y f) ->
inc (J x y) (disjoint_union f).
Proof.
move=> f x y; rewrite /disjoint_union /disjoint_union_fam; srw.
move=> yd xV; exists y; bw; split; fprops.
Qed.
Theorem disjoint_union_pr: forall f,
fgraph f -> exists x,
source x = disjoint_union f &
target x = unionb f &
surjection x &
(mutually_disjoint f -> bijection x).
Proof.
move=> f fgf.
set (aux:= (L (domain f) (fun i => product (V i f) (singleton i)))).
move: (partion_union_disjoint fgf) =>pf.
have da: (domain aux = domain f) by rewrite /aux; bw.
pose h i := BL P (product (V i f) (singleton i)) (unionb f).
have tap: (forall i, inc i (domain aux) ->
transf_axioms P (product (V i f) (singleton i)) (unionb f)).
by move => i id x; aw; move=> [_ [pv Qi]]; union_tac; ue.
have fpi: (forall i, inc i (domain aux) ->
function_prop (h i) (V i aux)(unionb f)).
move=> i id; split.
by apply bl_function; apply tap.
by rewrite /h/aux; aw; bw; auto; ue.
case (extension_partition pf fpi)=> [[g [fpg aag]] _].
move: aag; bw; rewrite/disjoint_union_fam; bw; move=> aag.
move: fpg=> [fg [sg tg]].
exists g; split=>//; split=>//.
have sug: (surjection g).
apply surjective_pr5=>//; rewrite /related.
rewrite tg sg; move=> y; srw; move=> [z [zdf yg]].
set (t:= J y z).
have tsg: inc t (disjoint_union f) by rewrite /t; apply inc_disjoint_union.
exists t; split=>//.
suff: y = W t g by move=>->; apply W_pr3 =>//; rewrite sg.
have ps: inc t (product (V z f) (singleton z)) by rewrite /t; aw; fprops.
move: (aag _ zdf); bw.
rewrite /agrees_on; move=> [_ [_ sv]]; move: (sv _ ps).
rewrite /h; aw.
by move=> ->; rewrite /t; aw.
by apply tap; ue.
by rewrite /t; aw; intuition.
split=>//.
move=> mf; split=>//; split=>//.
move=> x y; rewrite sg; move=> xsg ysg.
move: (du_index_pr xsg)=> [qxd [pV xp]].
move: (du_index_pr ysg)=> [qyd [pyV yp]].
move: (aag _ qxd); bw; move=> [ _ [_ ax]].
move: (aag _ qyd); bw; move=> [ _ [_ ay]].
have xs: inc x (product (V (Q x) f) (singleton (Q x))) by aw.
have ys: inc y (product (V (Q y) f) (singleton (Q y))) by aw.
rewrite (ax _ xs) (ay _ ys).
rewrite/h ;aw ; fprops; try apply tap; try rewrite da=>//.
move=> sp;move: (mf _ _ qxd qyd).
case=> sq; first by apply pair_exten.
red in sq.
by empty_tac1 (P x); apply intersection2_inc=>//; by rewrite sp.
Qed.
Bourbaki has an axiom (not needed here) that says that the powerset exists
Canonical extension of f:E->F to powerset E -> powerset F
Definition extension_to_parts f :=
BL(image_by_fun f) (powerset (source f)) (powerset (target f)).
Lemma etp_axioms: forall f,
is_correspondence f ->
transf_axioms (image_by_fun f) (powerset (source f))
(powerset (target f)).
Proof.
move=> f cf x; aw; move => xs; rewrite /image_by_fun.
apply sub_trans with (range (graph f)); fprops; apply sub_image_by_graph.
Qed.
Lemma etp_function: forall f,
is_correspondence f -> is_function (extension_to_parts f).
Proof.
rewrite /extension_to_parts=> f cf; apply bl_function;by apply etp_axioms.
Qed.
Lemma etp_W: forall f x,
is_correspondence f -> sub x (source f)
-> W x (extension_to_parts f) = image_by_fun f x.
Proof.
rewrite /extension_to_parts=> f x cf s; aw;by apply etp_axioms.
Qed.
Hint Resolve etp_function : fprops.
BL(image_by_fun f) (powerset (source f)) (powerset (target f)).
Lemma etp_axioms: forall f,
is_correspondence f ->
transf_axioms (image_by_fun f) (powerset (source f))
(powerset (target f)).
Proof.
move=> f cf x; aw; move => xs; rewrite /image_by_fun.
apply sub_trans with (range (graph f)); fprops; apply sub_image_by_graph.
Qed.
Lemma etp_function: forall f,
is_correspondence f -> is_function (extension_to_parts f).
Proof.
rewrite /extension_to_parts=> f cf; apply bl_function;by apply etp_axioms.
Qed.
Lemma etp_W: forall f x,
is_correspondence f -> sub x (source f)
-> W x (extension_to_parts f) = image_by_fun f x.
Proof.
rewrite /extension_to_parts=> f x cf s; aw;by apply etp_axioms.
Qed.
Hint Resolve etp_function : fprops.
Morphism properties of extension
Lemma etp_composable: forall f g,
composableC g f ->
(extension_to_parts g) \coP (extension_to_parts f).
Proof.
move=>f g [cf [cg cfg]]; red; ee; rewrite /extension_to_parts; aw; ue.
Qed.
Lemma etp_compose: forall f g,
composableC g f ->
(extension_to_parts g) \co (extension_to_parts f)
= extension_to_parts (g \co f).
Proof.
move=> f g cgf; move: (etp_composable cgf)=>cegf.
move: cgf=> [cf [cg cfg]].
have ccg: (is_correspondence (compose g f)) by apply compose_correspondence.
apply function_exten; try (fct_tac; fprops); fprops.
by rewrite /extension_to_parts; aw.
by rewrite /extension_to_parts; aw.
move=>x xs;awi xs.
have xs1: sub x (source f) by move: xs;rewrite /extension_to_parts; aw.
have xs2: sub x (source (g \co f)) by aw.
aw; repeat rewrite etp_W=>//; first by symmetry; apply compose_of_sets.
rewrite cfg /image_by_fun; apply sub_trans with (range (graph f)).
apply sub_image_by_graph.
fprops.
Qed.
Lemma etp_identity: forall x,
extension_to_parts (identity x) = identity (powerset x).
Proof.
move => x.
move:(@identity_corresp x)=> ic.
apply function_exten; fprops; try solve [rewrite/extension_to_parts; aw].
move=> y ys;
have Y: sub y (source (identity x)) by move:ys; rewrite/extension_to_parts; aw.
rewrite etp_W=>//.
move:Y; aw=> syx; srw; aw.
have fi: (is_function (identity x)) by fprops.
by set_extens t; aw; [move=> [u [uy]] <- | move=> ty; ex_tac]; srw; apply syx.
Qed.
Lemma composable_for_function: forall f g,
g \coP f -> composableC g f.
Proof. move=> f g [[ fg _] [[ff _ ] eq]]; hnf; intuition. Qed.
Theorem etp_surjective: forall f,
surjection f -> surjection (extension_to_parts f).
Proof.
move=> f sf; apply surj_if_exists_right_inv.
move: (exists_right_inv_from_surj sf)=> [r [cfr cfri]].
move:(composable_for_function cfr)=> cfr'.
exists (extension_to_parts r); split; first by apply etp_composable.
rewrite etp_compose // cfri etp_identity /extension_to_parts; aw.
Qed.
Theorem etp_injective: forall f,
injection f -> injection (extension_to_parts f).
Proof.
move=> f inf.
case(emptyset_dichot (source f))=> eq.
move: inf=> [[ cf _] _]; split; fprops.
rewrite {1 2} /extension_to_parts=> x y.
by rewrite eq powerset_emptyset; aw => -> ->; intuition.
case (exists_left_inv_from_inj inf eq) => [r [cfr cfri]].
move:(composable_for_function cfr)=> cfr'.
apply inj_if_exists_left_inv.
exists (extension_to_parts r); split; first by apply etp_composable.
rewrite etp_compose // cfri etp_identity /extension_to_parts; aw.
Qed.
composableC g f ->
(extension_to_parts g) \coP (extension_to_parts f).
Proof.
move=>f g [cf [cg cfg]]; red; ee; rewrite /extension_to_parts; aw; ue.
Qed.
Lemma etp_compose: forall f g,
composableC g f ->
(extension_to_parts g) \co (extension_to_parts f)
= extension_to_parts (g \co f).
Proof.
move=> f g cgf; move: (etp_composable cgf)=>cegf.
move: cgf=> [cf [cg cfg]].
have ccg: (is_correspondence (compose g f)) by apply compose_correspondence.
apply function_exten; try (fct_tac; fprops); fprops.
by rewrite /extension_to_parts; aw.
by rewrite /extension_to_parts; aw.
move=>x xs;awi xs.
have xs1: sub x (source f) by move: xs;rewrite /extension_to_parts; aw.
have xs2: sub x (source (g \co f)) by aw.
aw; repeat rewrite etp_W=>//; first by symmetry; apply compose_of_sets.
rewrite cfg /image_by_fun; apply sub_trans with (range (graph f)).
apply sub_image_by_graph.
fprops.
Qed.
Lemma etp_identity: forall x,
extension_to_parts (identity x) = identity (powerset x).
Proof.
move => x.
move:(@identity_corresp x)=> ic.
apply function_exten; fprops; try solve [rewrite/extension_to_parts; aw].
move=> y ys;
have Y: sub y (source (identity x)) by move:ys; rewrite/extension_to_parts; aw.
rewrite etp_W=>//.
move:Y; aw=> syx; srw; aw.
have fi: (is_function (identity x)) by fprops.
by set_extens t; aw; [move=> [u [uy]] <- | move=> ty; ex_tac]; srw; apply syx.
Qed.
Lemma composable_for_function: forall f g,
g \coP f -> composableC g f.
Proof. move=> f g [[ fg _] [[ff _ ] eq]]; hnf; intuition. Qed.
Theorem etp_surjective: forall f,
surjection f -> surjection (extension_to_parts f).
Proof.
move=> f sf; apply surj_if_exists_right_inv.
move: (exists_right_inv_from_surj sf)=> [r [cfr cfri]].
move:(composable_for_function cfr)=> cfr'.
exists (extension_to_parts r); split; first by apply etp_composable.
rewrite etp_compose // cfri etp_identity /extension_to_parts; aw.
Qed.
Theorem etp_injective: forall f,
injection f -> injection (extension_to_parts f).
Proof.
move=> f inf.
case(emptyset_dichot (source f))=> eq.
move: inf=> [[ cf _] _]; split; fprops.
rewrite {1 2} /extension_to_parts=> x y.
by rewrite eq powerset_emptyset; aw => -> ->; intuition.
case (exists_left_inv_from_inj inf eq) => [r [cfr cfri]].
move:(composable_for_function cfr)=> cfr'.
apply inj_if_exists_left_inv.
exists (extension_to_parts r); split; first by apply etp_composable.
rewrite etp_compose // cfri etp_identity /extension_to_parts; aw.
Qed.
The set of functional graphs is denoted by F^E, the set of functions by calF(E; F)
Definition set_of_functions x y :=
Zo (set_of_correspondences x y)
(fun z => fgraph (graph z) & x = domain (graph z)).
Definition set_of_permutations E :=
(Zo (set_of_functions E E)(fun z=> bijection z)).
Lemma set_of_functions_rw: forall x y f,
inc f (set_of_functions x y) =
(is_function f & source f = x & target f = y).
Proof.
rewrite /is_function/set_of_functions => x y f.
apply iff_eq; rewrite Z_rw set_of_correspondences_rw; intuition; ue.
Qed.
Hint Rewrite set_of_functions_rw : aw.
Lemma inc_set_of_gfunctions: forall f,
is_function f -> inc (graph f) (set_of_gfunctions (source f) (target f)).
Proof.
move=> f [[]]; rewrite /set_of_gfunctions Z_rw; move=> tf p1 p2; ee; aw.
Qed.
Lemma set_of_gfunctions_inc: forall x y z,
inc z (set_of_gfunctions x y) -> exists f,
is_function f & source f = x & target f = y & graph f =z.
Proof.
move=> x y z.
rewrite /set_of_gfunctions Z_rw; aw; rewrite corr_propcc.
move=> hh; exists (corresp x y z); aw.
intuition; apply is_function_pr=>//.
Qed.
Definition set_of_sub_functions (x y: Set) :=
unionf(powerset x)(fun z=> (set_of_functions z y)).
Lemma set_of_sub_functions_rw: forall x y f,
inc f (set_of_sub_functions x y) =
(is_function f & sub (source f) x & target f = y).
Proof.
move=> x y f; rewrite/set_of_sub_functions; apply iff_eq; srw.
by move=> [z [zp ff]]; awi ff; awi zp; intuition; ue.
by move=> [ff [ss tf]]; exists (source f); aw; intuition.
Qed.
Lemma empty_source_graph: forall f,
is_function f -> source f = emptyset -> graph f = emptyset.
Proof.
move => f ff sfe; empty_tac x xp.
elim (emptyset_pr (x:=(P x))); rewrite -sfe; graph_tac.
Qed.
Lemma empty_target_graph: forall f,
is_function f -> target f = emptyset -> graph f = emptyset.
Proof.
move => f ff sfe; empty_tac x xp.
elim (emptyset_pr (x:= W (P x)f)); rewrite -sfe; graph_tac; graph_tac.
Qed.
Lemma set_of_functions_extens: forall x y a b,
inc a (set_of_functions x y) -> inc b (set_of_functions x y) ->
graph a = graph b -> a = b.
Proof.
move=> x y a b; aw; move=> [fa [sa ta]] [fb [sb tb]] sg.
apply function_exten1=>//; ue.
Qed.
Lemma small_set_of_functions_source: forall y,
small_set (set_of_functions emptyset y).
Proof.
move=> y u v uf vf; apply (set_of_functions_extens uf vf).
move: uf vf; aw;move => [fa [sa _]] [fb [sb _]].
do 2 rewrite empty_source_graph=>//.
Qed.
Lemma small_set_of_functions_target: forall x,
small_set (set_of_functions x emptyset).
Proof.
move=> y u v uf vf; apply (set_of_functions_extens uf vf).
move: uf vf; aw;move => [fa [_ ta ]] [fb [_ tb ]].
do 2 rewrite empty_target_graph=>//.
Qed.
Lemma empty_set_of_functions_target: forall x y,
(x = emptyset \/ nonempty y) -> nonempty (set_of_functions x y).
Proof.
move=> x y; case=> p.
exists (BL (fun z:Set => z) x y).
aw; intuition; apply bl_function; rewrite p=> t tx; elim (emptyset_pr tx).
move: p => [p py]; exists (constant_function x y p py).
rewrite /constant_function; aw; fprops.
Qed.
Canonical isomorphism between calF(E,F) and F^E
Lemma graph_axioms: forall x y,
transf_axioms graph (set_of_functions x y) (set_of_gfunctions x y).
Proof. by move=>x y g; aw; move=> [fg []] <- <-; apply inc_set_of_gfunctions.
Qed.
Lemma graph_bijective: forall x y,
bijection (BL graph (set_of_functions x y) (set_of_gfunctions x y)).
Proof.
move=> x y; apply bl_bijective.
by apply graph_axioms.
by move=> u v us vs g; apply (set_of_functions_extens us vs g).
move=> z zg ; move: (set_of_gfunctions_inc zg) => [f pf].
by exists f; aw; intuition.
Qed.
Lemma set_of_functions_equipotent: forall x y,
(set_of_functions x y) \Eq (set_of_gfunctions x y).
Proof.
move=> x y; exists (BL graph (set_of_functions x y) (set_of_gfunctions x y)).
aw; intuition; apply graph_bijective.
Qed.
Isomorphism between calF(E,F) and calF(E',F') for equipotent sets
Definition compose3function u v :=
BL (fun f => (v \co f) \co u)
(set_of_functions (target u) (source v))
(set_of_functions (source u) (target v)).
Lemma c3f_axioms: forall u v,
is_function u -> is_function v ->
transf_axioms (fun f => compose (compose v f) u)
(set_of_functions (target u) (source v))
(set_of_functions (source u) (target v)).
Proof.
move=>u v fu fv x; aw; move => [fx [sx tx]]; intuition.
fct_tac; [fct_tac | aw].
Qed.
Lemma c3f_function: forall u v,
is_function u -> is_function v -> is_function (compose3function u v).
Proof.
rewrite /compose3function=> u v fu fv.
by apply bl_function; apply c3f_axioms.
Qed.
Lemma c3f_W: forall u v f,
is_function u -> is_function v ->
is_function f -> source f = target u -> target f = source v ->
W f (compose3function u v) = (v \co f) \co u.
Proof.
rewrite /compose3function=> u v f fu fv ff sf tf.
by aw; [apply c3f_axioms=>// | aw].
Qed.
Theorem c3f_injective: forall u v,
surjection u -> injection v -> injection (compose3function u v).
Proof.
rewrite /compose3function=> u v su iv.
have fu: is_function u by fct_tac.
have fv: is_function v by fct_tac.
apply bl_injective; first by apply c3f_axioms.
move=> f g.
case (emptyset_dichot (source v)) => sve.
move=> i1 i2 c.
by apply (small_set_of_functions_target (x:= target u)); ue.
move: (exists_right_inv_from_surj su) => [s [cus cusi]].
move: (exists_left_inv_from_inj iv sve).
aw; move=> [r [crv crvi]][ff [sf tf]][fg [sg tg]] ceq.
have cvf: v \coP f by hnf.
have cvg: v \coP g by hnf.
have fvf: is_function (v \co f) by fct_tac.
have fvg: is_function (v \co g) by fct_tac.
have cvfu: (v \co f) \coP u by hnf; aw.
have cvgu: (v \co g) \coP u by hnf; aw.
move:(f_equal (fun i => i \co s) ceq).
rewrite compose_assoc // (@compose_assoc (v \co g) u s) // cusi.
have s1: (source (v \co f) = target u) by aw.
have s2: (source (v \co g) = target u) by aw.
rewrite -{1} s1 -s2 compose_id_right // compose_id_right //; move=> eq.
move: (f_equal (fun i => r \co i) eq).
rewrite -compose_assoc // -compose_assoc // crvi -{1}tf -tg.
rewrite compose_id_left // compose_id_left //.
Qed.
Theorem c3f_surjective: forall u v,
(nonempty (source u) \/ (nonempty (source v)) \/ (nonempty (target v))
\/ target u = emptyset) ->
injection u -> surjection v -> surjection (compose3function u v).
Proof.
move=> u v nehyp iu sv.
have fu: is_function u by fct_tac.
have fv: is_function v by fct_tac.
move: (c3f_function fu fv)=> fc3.
set (c3f:= compose3function u v) in *.
have t3f:(target c3f= set_of_functions (source u) (target v))
by rewrite /compose3function bl_target.
have s3f :(source c3f =set_of_functions (target u) (source v) )
by rewrite /compose3function bl_source.
apply surjective_pr6=>//; rewrite t3f s3f => y ytc3.
case (emptyset_dichot (source u)) => nesu.
have p:(target u = emptyset \/ nonempty (source v)).
case nehyp.
by rewrite nesu=> h; elim (not_nonempty_empty h).
case; first (by auto); case; last by auto.
by move=> [w wt]; move: (surjective_pr sv wt)=> [x [xs _]];right; ex_tac.
have: (nonempty (set_of_functions (target u) (source v))).
by apply (empty_set_of_functions_target p).
rewrite -s3f; move=> [f fts].
ex_tac.
apply (small_set_of_functions_source (y:=target v)); rewrite - nesu=>//.
by rewrite -t3f; fprops.
move: (exists_right_inv_from_surj sv)=> [s [cvs cvsi]].
move: (exists_left_inv_from_inj iu nesu) => [r [cru crui]].
have Hb: is_function s by fct_tac.
have Hc: is_function r by fct_tac.
move:ytc3; aw; move=> [fy [sy ty]].
set (f:= (s \co y) \co r).
have sytr: source y = target r by rewrite sy; move: (f_equal target crui); aw.
have tytr: source s = target y by rewrite ty; move: (f_equal source cvsi); aw.
have csy: s \coP y by hnf; intuition.
have fcyr: is_function (y \co r) by fct_tac.
have csyr: s \coP (y \co r) by hnf; aw.
have tcyr : (target (y \co r) = target v) by aw.
have fcsy: is_function (s \co y) by fct_tac.
have ff: is_function f by rewrite/f; fct_tac=>//; aw.
have sf: source f = target u by move: cru=> [_[_]]; rewrite /f; aw.
have tf: target f = source v by move: cvs=> [_[_]]; rewrite /f; aw.
have fs: inc f (set_of_functions (target u) (source v)). by aw; auto.
ex_tac; rewrite c3f_W=>//.
rewrite/f (@compose_assoc s y r) // -(@compose_assoc v s (y \co r)) //.
rewrite cvsi -tcyr compose_id_left // compose_assoc //.
rewrite crui -sy compose_id_right //.
Qed.
Lemma c3f_bijective: forall u v,
bijection u -> bijection v -> bijection (compose3function u v).
Proof.
move=> u v [iu su][iv sv]; split.
by apply c3f_injective=>//.
apply c3f_surjective=>//.
case (emptyset_dichot (target u)); first by auto.
move=> [y yt]; move: (surjective_pr su yt)=> [x [xs _]]; left; ex_tac.
Qed.
BL (fun f => (v \co f) \co u)
(set_of_functions (target u) (source v))
(set_of_functions (source u) (target v)).
Lemma c3f_axioms: forall u v,
is_function u -> is_function v ->
transf_axioms (fun f => compose (compose v f) u)
(set_of_functions (target u) (source v))
(set_of_functions (source u) (target v)).
Proof.
move=>u v fu fv x; aw; move => [fx [sx tx]]; intuition.
fct_tac; [fct_tac | aw].
Qed.
Lemma c3f_function: forall u v,
is_function u -> is_function v -> is_function (compose3function u v).
Proof.
rewrite /compose3function=> u v fu fv.
by apply bl_function; apply c3f_axioms.
Qed.
Lemma c3f_W: forall u v f,
is_function u -> is_function v ->
is_function f -> source f = target u -> target f = source v ->
W f (compose3function u v) = (v \co f) \co u.
Proof.
rewrite /compose3function=> u v f fu fv ff sf tf.
by aw; [apply c3f_axioms=>// | aw].
Qed.
Theorem c3f_injective: forall u v,
surjection u -> injection v -> injection (compose3function u v).
Proof.
rewrite /compose3function=> u v su iv.
have fu: is_function u by fct_tac.
have fv: is_function v by fct_tac.
apply bl_injective; first by apply c3f_axioms.
move=> f g.
case (emptyset_dichot (source v)) => sve.
move=> i1 i2 c.
by apply (small_set_of_functions_target (x:= target u)); ue.
move: (exists_right_inv_from_surj su) => [s [cus cusi]].
move: (exists_left_inv_from_inj iv sve).
aw; move=> [r [crv crvi]][ff [sf tf]][fg [sg tg]] ceq.
have cvf: v \coP f by hnf.
have cvg: v \coP g by hnf.
have fvf: is_function (v \co f) by fct_tac.
have fvg: is_function (v \co g) by fct_tac.
have cvfu: (v \co f) \coP u by hnf; aw.
have cvgu: (v \co g) \coP u by hnf; aw.
move:(f_equal (fun i => i \co s) ceq).
rewrite compose_assoc // (@compose_assoc (v \co g) u s) // cusi.
have s1: (source (v \co f) = target u) by aw.
have s2: (source (v \co g) = target u) by aw.
rewrite -{1} s1 -s2 compose_id_right // compose_id_right //; move=> eq.
move: (f_equal (fun i => r \co i) eq).
rewrite -compose_assoc // -compose_assoc // crvi -{1}tf -tg.
rewrite compose_id_left // compose_id_left //.
Qed.
Theorem c3f_surjective: forall u v,
(nonempty (source u) \/ (nonempty (source v)) \/ (nonempty (target v))
\/ target u = emptyset) ->
injection u -> surjection v -> surjection (compose3function u v).
Proof.
move=> u v nehyp iu sv.
have fu: is_function u by fct_tac.
have fv: is_function v by fct_tac.
move: (c3f_function fu fv)=> fc3.
set (c3f:= compose3function u v) in *.
have t3f:(target c3f= set_of_functions (source u) (target v))
by rewrite /compose3function bl_target.
have s3f :(source c3f =set_of_functions (target u) (source v) )
by rewrite /compose3function bl_source.
apply surjective_pr6=>//; rewrite t3f s3f => y ytc3.
case (emptyset_dichot (source u)) => nesu.
have p:(target u = emptyset \/ nonempty (source v)).
case nehyp.
by rewrite nesu=> h; elim (not_nonempty_empty h).
case; first (by auto); case; last by auto.
by move=> [w wt]; move: (surjective_pr sv wt)=> [x [xs _]];right; ex_tac.
have: (nonempty (set_of_functions (target u) (source v))).
by apply (empty_set_of_functions_target p).
rewrite -s3f; move=> [f fts].
ex_tac.
apply (small_set_of_functions_source (y:=target v)); rewrite - nesu=>//.
by rewrite -t3f; fprops.
move: (exists_right_inv_from_surj sv)=> [s [cvs cvsi]].
move: (exists_left_inv_from_inj iu nesu) => [r [cru crui]].
have Hb: is_function s by fct_tac.
have Hc: is_function r by fct_tac.
move:ytc3; aw; move=> [fy [sy ty]].
set (f:= (s \co y) \co r).
have sytr: source y = target r by rewrite sy; move: (f_equal target crui); aw.
have tytr: source s = target y by rewrite ty; move: (f_equal source cvsi); aw.
have csy: s \coP y by hnf; intuition.
have fcyr: is_function (y \co r) by fct_tac.
have csyr: s \coP (y \co r) by hnf; aw.
have tcyr : (target (y \co r) = target v) by aw.
have fcsy: is_function (s \co y) by fct_tac.
have ff: is_function f by rewrite/f; fct_tac=>//; aw.
have sf: source f = target u by move: cru=> [_[_]]; rewrite /f; aw.
have tf: target f = source v by move: cvs=> [_[_]]; rewrite /f; aw.
have fs: inc f (set_of_functions (target u) (source v)). by aw; auto.
ex_tac; rewrite c3f_W=>//.
rewrite/f (@compose_assoc s y r) // -(@compose_assoc v s (y \co r)) //.
rewrite cvsi -tcyr compose_id_left // compose_assoc //.
rewrite crui -sy compose_id_right //.
Qed.
Lemma c3f_bijective: forall u v,
bijection u -> bijection v -> bijection (compose3function u v).
Proof.
move=> u v [iu su][iv sv]; split.
by apply c3f_injective=>//.
apply c3f_surjective=>//.
case (emptyset_dichot (target u)); first by auto.
move=> [y yt]; move: (surjective_pr su yt)=> [x [xs _]]; left; ex_tac.
Qed.
Isomorphism between calF(BxC,A) and calF(C, calF(B,A)) and calF(B, calF(C,A))
Definition partial_fun_axioms f :=
is_function f & exists a, exists b, source f = product a b.
Hint Rewrite product_domain product_range : dw.
Lemma partial_fun_axioms_pr : forall f,
partial_fun_axioms f ->
source f = product (domain (source f))(range (source f)).
Proof.
move=> f [ff [a [b sf]]]; rewrite sf.
case (emptyset_dichot a); first by move=> ->; srw.
case (emptyset_dichot b); first by move=> ->; srw.
by move=> h1 h2; dw.
Qed.
Definition first_partial_fun f y:=
BL(fun x => W (J x y) f) (domain (source f)) (target f).
Definition second_partial_fun f x:=
BL(fun y => W (J x y) f) (range (source f)) (target f).
Definition first_partial_function f:=
BL(fun y => first_partial_fun f y) (range (source f))
(set_of_functions (domain (source f)) (target f)).
Definition second_partial_function f:=
BL(fun x => second_partial_fun f x) (domain (source f))
(set_of_functions (range (source f)) (target f)).
Definition first_partial_map b c a:=
BL (fun f=> first_partial_function f)
(set_of_functions (product b c) a)
(set_of_functions c (set_of_functions b a)).
Definition second_partial_map b c a:=
BL (fun f=> second_partial_function f)
(set_of_functions (product b c) a)
(set_of_functions b (set_of_functions c a)).
Lemma fpf_axioms: forall f y,
partial_fun_axioms f -> inc y (range (source f)) ->
transf_axioms (fun x => W (J x y) f) (domain (source f)) (target f).
Proof.
move=> f y pfa; move: (partial_fun_axioms_pr pfa)=> pfb.
move=> yr x xs; move: pfa=>[ff _]; apply inc_W_target=> //; ue.
Qed.
Lemma spf_axioms: forall f x,
partial_fun_axioms f -> inc x (domain (source f)) ->
transf_axioms (fun y => W (J x y) f) (range (source f)) (target f).
Proof.
move=> f y pfa; move: (partial_fun_axioms_pr pfa)=> pfb.
move=> yr x xs; move: pfa=>[ff _]; apply inc_W_target=> //; ue.
Qed.
Lemma fpf_function :forall f y,
partial_fun_axioms f -> inc y (range (source f)) ->
is_function (first_partial_fun f y).
Proof.
rewrite/first_partial_fun=> f y pfa yr.
by apply bl_function; apply fpf_axioms.
Qed.
Lemma spf_function :forall f x,
partial_fun_axioms f -> inc x (domain (source f)) ->
is_function (second_partial_fun f x).
Proof.
rewrite/second_partial_fun=> f y pfa yr.
by apply bl_function; apply spf_axioms.
Qed.
Lemma fpf_W :forall f x y,
partial_fun_axioms f -> inc x (domain (source f)) ->
inc y (range (source f)) ->
W x (first_partial_fun f y) = W (J x y) f.
Proof.
rewrite/first_partial_fun=> f x y pfa xd yr.
by rewrite bl_W //; apply fpf_axioms.
Qed.
Lemma spf_W :forall f x y,
partial_fun_axioms f -> inc x (domain (source f)) ->
inc y (range (source f)) ->
W y (second_partial_fun f x) = W (J x y) f.
Proof.
rewrite/first_partial_fun=> f x y pfa xd yr.
by rewrite bl_W //; apply spf_axioms.
Qed.
Lemma fpfa_axioms: forall f,
partial_fun_axioms f ->
transf_axioms (fun y => (first_partial_fun f y))(range (source f))
(set_of_functions (domain (source f)) (target f)).
Proof.
move=>f pfa x xr; aw; split; first by apply fpf_function.
by rewrite /first_partial_fun; aw.
Qed.
Lemma spfa_axioms: forall f ,
partial_fun_axioms f ->
transf_axioms (fun x => second_partial_fun f x) (domain (source f))
(set_of_functions (range (source f)) (target f)).
Proof.
move=>f pfa x xr; aw; split; first by apply spf_function.
by rewrite /second_partial_fun; aw.
Qed.
Lemma fpfa_function: forall f,
partial_fun_axioms f -> is_function (first_partial_function f).
Proof.
rewrite /first_partial_function=> f fpa.
by apply bl_function; apply fpfa_axioms.
Qed.
Lemma spfa_function: forall f ,
partial_fun_axioms f -> is_function (second_partial_function f).
Proof.
rewrite /second_partial_function=> f fpa.
by apply bl_function; apply spfa_axioms.
Qed.
Lemma fpfa_W: forall f y,
partial_fun_axioms f -> inc y (range (source f)) ->
W y (first_partial_function f) = first_partial_fun f y.
Proof.
rewrite /first_partial_function=>f y pfa yr.
by rewrite bl_W //; apply fpfa_axioms.
Qed.
Lemma spfa_W: forall f x,
partial_fun_axioms f -> inc x (domain (source f)) ->
W x (second_partial_function f) = second_partial_fun f x.
Proof.
rewrite /second_partial_function=>f y pfa yr.
by rewrite bl_W //; apply spfa_axioms.
Qed.
Lemma fpfb_axioms: forall a b c,
nonempty b -> nonempty c ->
transf_axioms (fun f=> first_partial_function f)
(set_of_functions (product b c) a)
(set_of_functions c (set_of_functions b a)).
Proof.
move=>a b c neb nec x; aw; move=>[fx [sx tx]].
split.
by apply fpfa_function; hnf; split=>//; rewrite sx;exists b; exists c.
by rewrite/first_partial_function; aw ; rewrite sx tx; dw.
Qed.
Lemma spfb_axioms: forall a b c,
nonempty b -> nonempty c ->
transf_axioms (fun f=> second_partial_function f)
(set_of_functions (product b c) a)
(set_of_functions b (set_of_functions c a)).
Proof.
move=>a b c neb nec x; aw; move=>[fx [sx tx]].
split.
by apply spfa_function; hnf; split=>//; rewrite sx;exists b; exists c.
by rewrite/second_partial_function; aw ; rewrite sx tx; dw.
Qed.
Lemma fpfb_function: forall a b c,
nonempty a -> nonempty b -> is_function (first_partial_map a b c).
Proof.
rewrite /first_partial_map=> a b c nea neb.
by apply bl_function; apply fpfb_axioms.
Qed.
Lemma spfb_function: forall a b c,
nonempty a -> nonempty b -> is_function (second_partial_map a b c).
Proof.
rewrite /second_partial_map=> a b c nea neb.
by apply bl_function; apply spfb_axioms.
Qed.
Lemma fpfb_W: forall a b c f,
nonempty a -> nonempty b -> inc f (set_of_functions (product a b) c) ->
W f (first_partial_map a b c) = first_partial_function f.
Proof.
rewrite/first_partial_map=> a b c f nea neb fs; aw.
by apply fpfb_axioms.
Qed.
Lemma spfb_W: forall a b c f,
nonempty a -> nonempty b -> inc f (set_of_functions (product a b) c) ->
W f (second_partial_map a b c) = second_partial_function f.
Proof.
rewrite/second_partial_map=> a b c f nea neb fs; aw.
by apply spfb_axioms.
Qed.
Lemma fpfb_WW: forall a b c f x,
nonempty a -> nonempty b -> inc f (set_of_functions (product a b) c) ->
inc x (product a b) ->
W (P x) (W (Q x) (W f (first_partial_map a b c))) = W x f.
Proof.
move=> a b c f x nea neb fs xp; rewrite fpfb_W=>//.
move: fs; aw; move => [fs [sf tf]].
have pfaf: (partial_fun_axioms f) by split=>//; exists a; exists b.
move: xp; aw; move=>[xp [px qx]].
have qxr: (inc (Q x) (range (source f))) by rewrite sf; dw.
have pxr: (inc (P x) (domain (source f))) by rewrite sf; dw.
by rewrite fpfa_W // fpf_W // xp.
Qed.
Lemma spfb_WW: forall a b c f x,
nonempty a -> nonempty b -> inc f (set_of_functions (product a b) c) ->
inc x (product a b) ->
W (Q x) (W (P x) (W f (second_partial_map a b c))) = W x f.
Proof.
move=> a b c f x nea neb fs xp; rewrite spfb_W=>//.
move: fs; aw; move => [fs [sf tf]].
have pfaf: (partial_fun_axioms f) by split=>//; exists a; exists b.
move: xp; aw; move=>[xp [px qx]].
have qxr: (inc (Q x) (range (source f))) by rewrite sf; dw.
have pxr: (inc (P x) (domain (source f))) by rewrite sf; dw.
by rewrite spfa_W // spf_W // xp.
Qed.
Theorem fpfa_bijective: forall a b c,
nonempty a -> nonempty b -> bijection (first_partial_map a b c).
Proof.
move=> a b c nea neb.
set f:= _ a b c.
have sf:source f = set_of_functions (product a b) c.
by rewrite /f /first_partial_map; aw.
have tf: target f = set_of_functions b (set_of_functions a c)
by rewrite /f /first_partial_map; aw.
have ff: is_function f by apply (fpfb_function _ nea neb).
split.
split=>//.
move=> x y; rewrite sf=> xs ys eq.
have sx: source x = product a b by move: xs; aw; intuition.
have sy: source y = product a b by move: ys; aw; intuition.
have ha: forall z, inc z (source x) -> W z x = W z y.
rewrite sx => z zp.
by rewrite -(fpfb_WW nea neb xs zp) -(fpfb_WW nea neb ys zp) eq.
move: xs ys; aw; move=> [fx [_ tx]] [fy [_ ty]].
by apply function_exten=>//; try ue.
apply surjective_pr6=>//; move=> y yt.
set (g:= BL (fun x => W (P x) (W (Q x) y)) (product a b) c).
have ta: transf_axioms (fun x => W (P x) (W (Q x) y)) (product a b) c.
move=> x; aw; move=> [xp [pa qb]].
move: yt; rewrite tf; aw; move=> [fy [sy ty]].
have Wt: (inc (W (Q x) y) (target y)) by apply inc_W_target=>//; ue.
by move: Wt; rewrite ty; aw; move=> [fW [sW]] <-; apply inc_W_target=>//; ue.
have gsf: (inc g (source f)).
by rewrite /g sf; aw; intuition; apply bl_function.
ex_tac.
have: inc (W g f) (target f) by apply inc_W_target.
move:yt; rewrite tf; aw; move=> [fy [sy ty]] [fW [sW tw]].
apply function_exten=>//; try ue; move => x; rewrite sW=> xs.
have: inc (W x (W g f)) (target (W g f)) by apply inc_W_target=>//; ue.
have: inc (W x y) (target y) by apply inc_W_target=>//; ue.
rewrite tw ty; aw; move=> [fy1 [sy1 ty1]] [fW1 [sW1 tw1]].
apply function_exten=>//; try ue; rewrite sW1=> z zs.
have psg: (inc (J z x) (source g)) by rewrite /g; aw; fprops.
rewrite/g bl_source in psg; rewrite sf in gsf.
move: (fpfb_WW nea neb gsf psg); aw; move=> ->.
by rewrite /g bl_W; first by aw.
Qed.
Theorem spfa_bijective: forall a b c,
nonempty a -> nonempty b -> bijection (second_partial_map a b c).
Proof.
move=> a b c nea neb.
set (f:= second_partial_map a b c).
have sf:source f = set_of_functions (product a b) c
by rewrite /f /second_partial_map; aw.
have tf: target f = set_of_functions a (set_of_functions b c).
by rewrite /f /second_partial_map; aw.
have ff: is_function f by apply (spfb_function _ nea neb).
split.
split=>//.
move=> x y; rewrite sf=> xs ys eq.
have sx: source x = product a b by move: xs; aw; intuition.
have sy: source y = product a b by move: ys; aw; intuition.
have ha: forall z, inc z (source x) -> W z x = W z y.
rewrite sx => z zp.
by rewrite -(spfb_WW nea neb xs zp) -(spfb_WW nea neb ys zp) eq.
move: xs ys; aw; move=> [fx [_ tx]] [fy [_ ty]].
by apply function_exten=>//; try ue.
apply surjective_pr6=>//; move=> y yt.
set (g:= BL (fun x => W (Q x) (W (P x) y)) (product a b) c).
have ta: transf_axioms (fun x => W (Q x) (W (P x) y)) (product a b) c.
move=> x; aw; move=> [xp [pa qb]].
move: yt; rewrite tf; aw; move=> [fy [sy ty]].
have Wt: (inc (W (P x) y) (target y)) by apply inc_W_target=>//; ue.
by move: Wt; rewrite ty; aw; move=> [fW [sW]] <-; apply inc_W_target=>//; ue.
have gsf: (inc g (source f)).
by rewrite /g sf; aw; intuition; apply bl_function.
ex_tac.
have: inc (W g f) (target f) by apply inc_W_target.
move:yt; rewrite tf; aw; move=> [fy [sy ty]] [fW [sW tw]].
apply function_exten=>//; try ue; move => x; rewrite sW=> xs.
have: inc (W x (W g f)) (target (W g f)) by apply inc_W_target=>//; ue.
have :inc (W x y) (target y) by apply inc_W_target=>//; ue.
rewrite tw ty; aw; move=> [fy1 [sy1 ty1]] [fW1 [sW1 tw1]].
apply function_exten=>//; try ue; rewrite sW1=> z zs.
have psg: (inc (J x z) (source g)) by rewrite /g; aw; fprops.
rewrite/g bl_source in psg; rewrite sf in gsf.
move: (spfb_WW nea neb gsf psg); aw; move=> ->.
by rewrite /g bl_W; first by aw.
Qed.
End Bunion.
Export Bunion.
Module Bproduct.
is_function f & exists a, exists b, source f = product a b.
Hint Rewrite product_domain product_range : dw.
Lemma partial_fun_axioms_pr : forall f,
partial_fun_axioms f ->
source f = product (domain (source f))(range (source f)).
Proof.
move=> f [ff [a [b sf]]]; rewrite sf.
case (emptyset_dichot a); first by move=> ->; srw.
case (emptyset_dichot b); first by move=> ->; srw.
by move=> h1 h2; dw.
Qed.
Definition first_partial_fun f y:=
BL(fun x => W (J x y) f) (domain (source f)) (target f).
Definition second_partial_fun f x:=
BL(fun y => W (J x y) f) (range (source f)) (target f).
Definition first_partial_function f:=
BL(fun y => first_partial_fun f y) (range (source f))
(set_of_functions (domain (source f)) (target f)).
Definition second_partial_function f:=
BL(fun x => second_partial_fun f x) (domain (source f))
(set_of_functions (range (source f)) (target f)).
Definition first_partial_map b c a:=
BL (fun f=> first_partial_function f)
(set_of_functions (product b c) a)
(set_of_functions c (set_of_functions b a)).
Definition second_partial_map b c a:=
BL (fun f=> second_partial_function f)
(set_of_functions (product b c) a)
(set_of_functions b (set_of_functions c a)).
Lemma fpf_axioms: forall f y,
partial_fun_axioms f -> inc y (range (source f)) ->
transf_axioms (fun x => W (J x y) f) (domain (source f)) (target f).
Proof.
move=> f y pfa; move: (partial_fun_axioms_pr pfa)=> pfb.
move=> yr x xs; move: pfa=>[ff _]; apply inc_W_target=> //; ue.
Qed.
Lemma spf_axioms: forall f x,
partial_fun_axioms f -> inc x (domain (source f)) ->
transf_axioms (fun y => W (J x y) f) (range (source f)) (target f).
Proof.
move=> f y pfa; move: (partial_fun_axioms_pr pfa)=> pfb.
move=> yr x xs; move: pfa=>[ff _]; apply inc_W_target=> //; ue.
Qed.
Lemma fpf_function :forall f y,
partial_fun_axioms f -> inc y (range (source f)) ->
is_function (first_partial_fun f y).
Proof.
rewrite/first_partial_fun=> f y pfa yr.
by apply bl_function; apply fpf_axioms.
Qed.
Lemma spf_function :forall f x,
partial_fun_axioms f -> inc x (domain (source f)) ->
is_function (second_partial_fun f x).
Proof.
rewrite/second_partial_fun=> f y pfa yr.
by apply bl_function; apply spf_axioms.
Qed.
Lemma fpf_W :forall f x y,
partial_fun_axioms f -> inc x (domain (source f)) ->
inc y (range (source f)) ->
W x (first_partial_fun f y) = W (J x y) f.
Proof.
rewrite/first_partial_fun=> f x y pfa xd yr.
by rewrite bl_W //; apply fpf_axioms.
Qed.
Lemma spf_W :forall f x y,
partial_fun_axioms f -> inc x (domain (source f)) ->
inc y (range (source f)) ->
W y (second_partial_fun f x) = W (J x y) f.
Proof.
rewrite/first_partial_fun=> f x y pfa xd yr.
by rewrite bl_W //; apply spf_axioms.
Qed.
Lemma fpfa_axioms: forall f,
partial_fun_axioms f ->
transf_axioms (fun y => (first_partial_fun f y))(range (source f))
(set_of_functions (domain (source f)) (target f)).
Proof.
move=>f pfa x xr; aw; split; first by apply fpf_function.
by rewrite /first_partial_fun; aw.
Qed.
Lemma spfa_axioms: forall f ,
partial_fun_axioms f ->
transf_axioms (fun x => second_partial_fun f x) (domain (source f))
(set_of_functions (range (source f)) (target f)).
Proof.
move=>f pfa x xr; aw; split; first by apply spf_function.
by rewrite /second_partial_fun; aw.
Qed.
Lemma fpfa_function: forall f,
partial_fun_axioms f -> is_function (first_partial_function f).
Proof.
rewrite /first_partial_function=> f fpa.
by apply bl_function; apply fpfa_axioms.
Qed.
Lemma spfa_function: forall f ,
partial_fun_axioms f -> is_function (second_partial_function f).
Proof.
rewrite /second_partial_function=> f fpa.
by apply bl_function; apply spfa_axioms.
Qed.
Lemma fpfa_W: forall f y,
partial_fun_axioms f -> inc y (range (source f)) ->
W y (first_partial_function f) = first_partial_fun f y.
Proof.
rewrite /first_partial_function=>f y pfa yr.
by rewrite bl_W //; apply fpfa_axioms.
Qed.
Lemma spfa_W: forall f x,
partial_fun_axioms f -> inc x (domain (source f)) ->
W x (second_partial_function f) = second_partial_fun f x.
Proof.
rewrite /second_partial_function=>f y pfa yr.
by rewrite bl_W //; apply spfa_axioms.
Qed.
Lemma fpfb_axioms: forall a b c,
nonempty b -> nonempty c ->
transf_axioms (fun f=> first_partial_function f)
(set_of_functions (product b c) a)
(set_of_functions c (set_of_functions b a)).
Proof.
move=>a b c neb nec x; aw; move=>[fx [sx tx]].
split.
by apply fpfa_function; hnf; split=>//; rewrite sx;exists b; exists c.
by rewrite/first_partial_function; aw ; rewrite sx tx; dw.
Qed.
Lemma spfb_axioms: forall a b c,
nonempty b -> nonempty c ->
transf_axioms (fun f=> second_partial_function f)
(set_of_functions (product b c) a)
(set_of_functions b (set_of_functions c a)).
Proof.
move=>a b c neb nec x; aw; move=>[fx [sx tx]].
split.
by apply spfa_function; hnf; split=>//; rewrite sx;exists b; exists c.
by rewrite/second_partial_function; aw ; rewrite sx tx; dw.
Qed.
Lemma fpfb_function: forall a b c,
nonempty a -> nonempty b -> is_function (first_partial_map a b c).
Proof.
rewrite /first_partial_map=> a b c nea neb.
by apply bl_function; apply fpfb_axioms.
Qed.
Lemma spfb_function: forall a b c,
nonempty a -> nonempty b -> is_function (second_partial_map a b c).
Proof.
rewrite /second_partial_map=> a b c nea neb.
by apply bl_function; apply spfb_axioms.
Qed.
Lemma fpfb_W: forall a b c f,
nonempty a -> nonempty b -> inc f (set_of_functions (product a b) c) ->
W f (first_partial_map a b c) = first_partial_function f.
Proof.
rewrite/first_partial_map=> a b c f nea neb fs; aw.
by apply fpfb_axioms.
Qed.
Lemma spfb_W: forall a b c f,
nonempty a -> nonempty b -> inc f (set_of_functions (product a b) c) ->
W f (second_partial_map a b c) = second_partial_function f.
Proof.
rewrite/second_partial_map=> a b c f nea neb fs; aw.
by apply spfb_axioms.
Qed.
Lemma fpfb_WW: forall a b c f x,
nonempty a -> nonempty b -> inc f (set_of_functions (product a b) c) ->
inc x (product a b) ->
W (P x) (W (Q x) (W f (first_partial_map a b c))) = W x f.
Proof.
move=> a b c f x nea neb fs xp; rewrite fpfb_W=>//.
move: fs; aw; move => [fs [sf tf]].
have pfaf: (partial_fun_axioms f) by split=>//; exists a; exists b.
move: xp; aw; move=>[xp [px qx]].
have qxr: (inc (Q x) (range (source f))) by rewrite sf; dw.
have pxr: (inc (P x) (domain (source f))) by rewrite sf; dw.
by rewrite fpfa_W // fpf_W // xp.
Qed.
Lemma spfb_WW: forall a b c f x,
nonempty a -> nonempty b -> inc f (set_of_functions (product a b) c) ->
inc x (product a b) ->
W (Q x) (W (P x) (W f (second_partial_map a b c))) = W x f.
Proof.
move=> a b c f x nea neb fs xp; rewrite spfb_W=>//.
move: fs; aw; move => [fs [sf tf]].
have pfaf: (partial_fun_axioms f) by split=>//; exists a; exists b.
move: xp; aw; move=>[xp [px qx]].
have qxr: (inc (Q x) (range (source f))) by rewrite sf; dw.
have pxr: (inc (P x) (domain (source f))) by rewrite sf; dw.
by rewrite spfa_W // spf_W // xp.
Qed.
Theorem fpfa_bijective: forall a b c,
nonempty a -> nonempty b -> bijection (first_partial_map a b c).
Proof.
move=> a b c nea neb.
set f:= _ a b c.
have sf:source f = set_of_functions (product a b) c.
by rewrite /f /first_partial_map; aw.
have tf: target f = set_of_functions b (set_of_functions a c)
by rewrite /f /first_partial_map; aw.
have ff: is_function f by apply (fpfb_function _ nea neb).
split.
split=>//.
move=> x y; rewrite sf=> xs ys eq.
have sx: source x = product a b by move: xs; aw; intuition.
have sy: source y = product a b by move: ys; aw; intuition.
have ha: forall z, inc z (source x) -> W z x = W z y.
rewrite sx => z zp.
by rewrite -(fpfb_WW nea neb xs zp) -(fpfb_WW nea neb ys zp) eq.
move: xs ys; aw; move=> [fx [_ tx]] [fy [_ ty]].
by apply function_exten=>//; try ue.
apply surjective_pr6=>//; move=> y yt.
set (g:= BL (fun x => W (P x) (W (Q x) y)) (product a b) c).
have ta: transf_axioms (fun x => W (P x) (W (Q x) y)) (product a b) c.
move=> x; aw; move=> [xp [pa qb]].
move: yt; rewrite tf; aw; move=> [fy [sy ty]].
have Wt: (inc (W (Q x) y) (target y)) by apply inc_W_target=>//; ue.
by move: Wt; rewrite ty; aw; move=> [fW [sW]] <-; apply inc_W_target=>//; ue.
have gsf: (inc g (source f)).
by rewrite /g sf; aw; intuition; apply bl_function.
ex_tac.
have: inc (W g f) (target f) by apply inc_W_target.
move:yt; rewrite tf; aw; move=> [fy [sy ty]] [fW [sW tw]].
apply function_exten=>//; try ue; move => x; rewrite sW=> xs.
have: inc (W x (W g f)) (target (W g f)) by apply inc_W_target=>//; ue.
have: inc (W x y) (target y) by apply inc_W_target=>//; ue.
rewrite tw ty; aw; move=> [fy1 [sy1 ty1]] [fW1 [sW1 tw1]].
apply function_exten=>//; try ue; rewrite sW1=> z zs.
have psg: (inc (J z x) (source g)) by rewrite /g; aw; fprops.
rewrite/g bl_source in psg; rewrite sf in gsf.
move: (fpfb_WW nea neb gsf psg); aw; move=> ->.
by rewrite /g bl_W; first by aw.
Qed.
Theorem spfa_bijective: forall a b c,
nonempty a -> nonempty b -> bijection (second_partial_map a b c).
Proof.
move=> a b c nea neb.
set (f:= second_partial_map a b c).
have sf:source f = set_of_functions (product a b) c
by rewrite /f /second_partial_map; aw.
have tf: target f = set_of_functions a (set_of_functions b c).
by rewrite /f /second_partial_map; aw.
have ff: is_function f by apply (spfb_function _ nea neb).
split.
split=>//.
move=> x y; rewrite sf=> xs ys eq.
have sx: source x = product a b by move: xs; aw; intuition.
have sy: source y = product a b by move: ys; aw; intuition.
have ha: forall z, inc z (source x) -> W z x = W z y.
rewrite sx => z zp.
by rewrite -(spfb_WW nea neb xs zp) -(spfb_WW nea neb ys zp) eq.
move: xs ys; aw; move=> [fx [_ tx]] [fy [_ ty]].
by apply function_exten=>//; try ue.
apply surjective_pr6=>//; move=> y yt.
set (g:= BL (fun x => W (Q x) (W (P x) y)) (product a b) c).
have ta: transf_axioms (fun x => W (Q x) (W (P x) y)) (product a b) c.
move=> x; aw; move=> [xp [pa qb]].
move: yt; rewrite tf; aw; move=> [fy [sy ty]].
have Wt: (inc (W (P x) y) (target y)) by apply inc_W_target=>//; ue.
by move: Wt; rewrite ty; aw; move=> [fW [sW]] <-; apply inc_W_target=>//; ue.
have gsf: (inc g (source f)).
by rewrite /g sf; aw; intuition; apply bl_function.
ex_tac.
have: inc (W g f) (target f) by apply inc_W_target.
move:yt; rewrite tf; aw; move=> [fy [sy ty]] [fW [sW tw]].
apply function_exten=>//; try ue; move => x; rewrite sW=> xs.
have: inc (W x (W g f)) (target (W g f)) by apply inc_W_target=>//; ue.
have :inc (W x y) (target y) by apply inc_W_target=>//; ue.
rewrite tw ty; aw; move=> [fy1 [sy1 ty1]] [fW1 [sW1 tw1]].
apply function_exten=>//; try ue; rewrite sW1=> z zs.
have psg: (inc (J x z) (source g)) by rewrite /g; aw; fprops.
rewrite/g bl_source in psg; rewrite sf in gsf.
move: (spfb_WW nea neb gsf psg); aw; move=> ->.
by rewrite /g bl_W; first by aw.
Qed.
End Bunion.
Export Bunion.
Module Bproduct.
We may need the graph of a function of type aSet->
Definition gbcreate (a:Set) (f:a->Set) := (IM (fun y:a => J (Ro y) (f y))).
Lemma gbcreate_rw: forall (a:Set) (f:a->Set) x,
inc x (gbcreate f) = exists y:a, J (Ro y) (f y) =x.
Proof.
by rewrite/gbcreate=> a f x; apply iff_eq; rewrite IM_rw.
Qed.
Lemma gbcreate_graph: forall (a:Set) (f:a-> Set), is_graph (gbcreate f).
Proof.
move=> a f x; rewrite gbcreate_rw; move=> [y] <-; fprops.
Qed.
Lemma gbcreate_fgraph: forall (a:Set) (f:a-> Set), fgraph (gbcreate f).
Proof.
move=> a f; split; first by apply gbcreate_graph.
move=> x y; rewrite 2! gbcreate_rw.
move=> [u] <- [v] <-; aw => eq.
by rewrite (R_inj eq).
Qed.
Lemma gbcreate_domain : forall (a:Set) (f:a-> Set),
domain (gbcreate f) = a.
Proof.
move=> a f; set_extens t; dw; try apply gbcreate_graph.
move=> [y]; rewrite gbcreate_rw; move=> [z eq].
by rewrite - (pr1_def eq); apply R_inc.
move=> ta; exists (f (Bo ta)); rewrite gbcreate_rw.
by exists (Bo ta); rewrite (B_eq ta).
Qed.
Lemma gbcreate_V : forall (a:Set) (f:a-> Set) (x:a),
V (Ro x) (gbcreate f) = f x.
Proof.
move=> a f x.
have Jg: (inc (J (Ro x) (f x)) (gbcreate f)) by rewrite gbcreate_rw; exists x.
by move: (pr2_V (gbcreate_fgraph f) Jg); aw.
Qed.
Lemma gbcreate_rw: forall (a:Set) (f:a->Set) x,
inc x (gbcreate f) = exists y:a, J (Ro y) (f y) =x.
Proof.
by rewrite/gbcreate=> a f x; apply iff_eq; rewrite IM_rw.
Qed.
Lemma gbcreate_graph: forall (a:Set) (f:a-> Set), is_graph (gbcreate f).
Proof.
move=> a f x; rewrite gbcreate_rw; move=> [y] <-; fprops.
Qed.
Lemma gbcreate_fgraph: forall (a:Set) (f:a-> Set), fgraph (gbcreate f).
Proof.
move=> a f; split; first by apply gbcreate_graph.
move=> x y; rewrite 2! gbcreate_rw.
move=> [u] <- [v] <-; aw => eq.
by rewrite (R_inj eq).
Qed.
Lemma gbcreate_domain : forall (a:Set) (f:a-> Set),
domain (gbcreate f) = a.
Proof.
move=> a f; set_extens t; dw; try apply gbcreate_graph.
move=> [y]; rewrite gbcreate_rw; move=> [z eq].
by rewrite - (pr1_def eq); apply R_inc.
move=> ta; exists (f (Bo ta)); rewrite gbcreate_rw.
by exists (Bo ta); rewrite (B_eq ta).
Qed.
Lemma gbcreate_V : forall (a:Set) (f:a-> Set) (x:a),
V (Ro x) (gbcreate f) = f x.
Proof.
move=> a f x.
have Jg: (inc (J (Ro x) (f x)) (gbcreate f)) by rewrite gbcreate_rw; exists x.
by move: (pr2_V (gbcreate_fgraph f) Jg); aw.
Qed.
We have three equivalent definitions of the product; in any case,
an element of the product is a functional graph with fixed domain
Definition productb f:=
Zo (powerset (product (domain f) (unionb f)))
(fun z => fgraph z & domain z = domain f
& forall i, inc i (domain f) -> inc (V i z) (V i f)).
Definition productt (I:Set) (f:I->Set):= productb (gbcreate f).
Definition productf sf f:= productb (L sf f).
Lemma productb_rw: forall f x,
fgraph f ->
inc x (productb f) =
(fgraph x & domain x = domain f &
forall i, inc i (domain x) -> inc (V i x) (V i f)).
Proof.
rewrite /productb => f x fgf; rewrite Z_rw; apply iff_eq.
by move=> [h1 [h2 [h3 h4]]]; rewrite h3.
move=> [h2 [h3 h4]]; rewrite -h3; intuition.
aw; move=> t tx.
have gx: (is_graph x) by fprops.
move: (gx _ tx); aw; dw;move => px; split=>//; split.
by exists (Q t); rewrite px.
have pd: (inc (P t) (domain f)) by rewrite -h3;apply inc_pr1_domain.
by union_tac; rewrite (pr2_V h2 tx); apply h4; ue.
Qed.
Lemma productt_rw: forall (I:Set) (f:I->Set) x,
inc x (productt f) =
(fgraph x & domain x = I & forall i:I, inc (V (Ro i) x) (f i)).
Proof.
move=> I f x; rewrite /productt productb_rw; last by apply gbcreate_fgraph.
rewrite gbcreate_domain.
apply iff_eq; move=> [fgx [dx h]]; split=>//; split=>//.
move=> i; rewrite dx in h.
by move: (h _ (R_inc i)); rewrite gbcreate_V.
rewrite dx => i id; rewrite - (B_eq id) gbcreate_V.
by apply h.
Qed.
Lemma productf_rw: forall sf f x,
inc x (productf sf f) =
(fgraph x & domain x = sf &
forall i, inc i (domain x) -> inc (V i x) (f i)).
Proof.
move=> sf f x; rewrite /productf productb_rw; gprops; bw.
apply iff_eq; move=> [? [? h]]; split=>//; split=>//; move=> i id;
move: (h _ id); bw; ue.
Qed.
Hint Rewrite productf_rw productt_rw productb_rw: aw.
Lemma productt_exten: forall (I:Set) (f:I->Set) x x',
inc x (productt f) -> inc x' (productt f) ->
(forall i:I, V (Ro i) x = V (Ro i) x') ->
x = x'.
Proof.
move=> I f x x'; aw; move=> [fx [dx px]][fx' [dx' px']] sv.
apply fgraph_exten=>//; try ue.
by move=>v; rewrite dx=> vd; rewrite - (B_eq vd); apply sv.
Qed.
Lemma productb_exten: forall f x x',
fgraph f ->
inc x (productb f) -> inc x' (productb f) ->
(forall i, inc i (domain f) -> V i x = V i x') ->
x = x'.
Proof.
move=> f x x' H; aw; move=> [fx [dx px]][fx' [dx' px']] sv.
apply fgraph_exten=>//; ue.
Qed.
Lemma productf_exten: forall sf f x x',
inc x (productf sf f) -> inc x' (productf sf f) ->
(forall i, inc i sf -> V i x = V i x') ->
x = x'.
Proof.
rewrite/ productf=> sf f x x' xp xp' sv.
have fg: (fgraph (L sf f)) by gprops.
apply (productb_exten fg xp xp'); bw.
Qed.
Lemma trivial_fgraph: forall f, L emptyset f = emptyset.
Proof.
rewrite /L=> f; empty_tac x xe; move: xe; aw.
move=> [z [ze]]; elim (emptyset_pr ze).
Qed.
Zo (powerset (product (domain f) (unionb f)))
(fun z => fgraph z & domain z = domain f
& forall i, inc i (domain f) -> inc (V i z) (V i f)).
Definition productt (I:Set) (f:I->Set):= productb (gbcreate f).
Definition productf sf f:= productb (L sf f).
Lemma productb_rw: forall f x,
fgraph f ->
inc x (productb f) =
(fgraph x & domain x = domain f &
forall i, inc i (domain x) -> inc (V i x) (V i f)).
Proof.
rewrite /productb => f x fgf; rewrite Z_rw; apply iff_eq.
by move=> [h1 [h2 [h3 h4]]]; rewrite h3.
move=> [h2 [h3 h4]]; rewrite -h3; intuition.
aw; move=> t tx.
have gx: (is_graph x) by fprops.
move: (gx _ tx); aw; dw;move => px; split=>//; split.
by exists (Q t); rewrite px.
have pd: (inc (P t) (domain f)) by rewrite -h3;apply inc_pr1_domain.
by union_tac; rewrite (pr2_V h2 tx); apply h4; ue.
Qed.
Lemma productt_rw: forall (I:Set) (f:I->Set) x,
inc x (productt f) =
(fgraph x & domain x = I & forall i:I, inc (V (Ro i) x) (f i)).
Proof.
move=> I f x; rewrite /productt productb_rw; last by apply gbcreate_fgraph.
rewrite gbcreate_domain.
apply iff_eq; move=> [fgx [dx h]]; split=>//; split=>//.
move=> i; rewrite dx in h.
by move: (h _ (R_inc i)); rewrite gbcreate_V.
rewrite dx => i id; rewrite - (B_eq id) gbcreate_V.
by apply h.
Qed.
Lemma productf_rw: forall sf f x,
inc x (productf sf f) =
(fgraph x & domain x = sf &
forall i, inc i (domain x) -> inc (V i x) (f i)).
Proof.
move=> sf f x; rewrite /productf productb_rw; gprops; bw.
apply iff_eq; move=> [? [? h]]; split=>//; split=>//; move=> i id;
move: (h _ id); bw; ue.
Qed.
Hint Rewrite productf_rw productt_rw productb_rw: aw.
Lemma productt_exten: forall (I:Set) (f:I->Set) x x',
inc x (productt f) -> inc x' (productt f) ->
(forall i:I, V (Ro i) x = V (Ro i) x') ->
x = x'.
Proof.
move=> I f x x'; aw; move=> [fx [dx px]][fx' [dx' px']] sv.
apply fgraph_exten=>//; try ue.
by move=>v; rewrite dx=> vd; rewrite - (B_eq vd); apply sv.
Qed.
Lemma productb_exten: forall f x x',
fgraph f ->
inc x (productb f) -> inc x' (productb f) ->
(forall i, inc i (domain f) -> V i x = V i x') ->
x = x'.
Proof.
move=> f x x' H; aw; move=> [fx [dx px]][fx' [dx' px']] sv.
apply fgraph_exten=>//; ue.
Qed.
Lemma productf_exten: forall sf f x x',
inc x (productf sf f) -> inc x' (productf sf f) ->
(forall i, inc i sf -> V i x = V i x') ->
x = x'.
Proof.
rewrite/ productf=> sf f x x' xp xp' sv.
have fg: (fgraph (L sf f)) by gprops.
apply (productb_exten fg xp xp'); bw.
Qed.
Lemma trivial_fgraph: forall f, L emptyset f = emptyset.
Proof.
rewrite /L=> f; empty_tac x xe; move: xe; aw.
move=> [z [ze]]; elim (emptyset_pr ze).
Qed.
Projection function to component i
Definition pr_i f i:= BL(V i) (productb f) (V i f).
Definition pr_it (I:Set) (f:I ->Set):= pr_i (gbcreate f).
Lemma pri_axioms: forall f i,
fgraph f -> inc i (domain f) ->
transf_axioms (V i)(productb f)(V i f).
Proof. by rewrite /pr_i=> f i fg id x; aw; move=> [fx [dx]]; apply; ue.
Qed.
Lemma pri_function: forall f i,
fgraph f -> inc i (domain f) ->
is_function (pr_i f i).
Proof. by rewrite/pr_i=> f i fg id; apply bl_function; apply pri_axioms.
Qed.
Lemma pri_W: forall f i x,
fgraph f -> inc i (domain f) -> inc x (productb f) ->
W x (pr_i f i) = V i x.
Proof. by rewrite/pr_i=> f i x fg id xp; aw; apply pri_axioms.
Qed.
Hint Resolve pri_function : fprops.
Special case where the domain is empty
Lemma product_trivial :
productb emptyset = singleton emptyset.
Proof.
move: emptyset_fgraph=> h.
set_extens t; aw; srw.
by move=> [fgt]; rewrite empty_graph1; fprops; intuition.
move=> ->; rewrite emptyset_domain; intuition; empty_tac0.
Qed.
Lemma product2_trivial: forall f,
productb (L emptyset f) = singleton emptyset.
Proof. move=> f; rewrite trivial_fgraph; apply product_trivial.
Qed.
a product is empty iff all factors are empty
Lemma product_nonempty: forall f,
fgraph f -> (forall i, inc i (domain f) -> nonempty (V i f)) ->
nonempty (productb f).
Proof.
move=> f fg hyp.
exists (L(domain f)(fun i=> rep (V i f))).
aw; split; gprops; split; bw.
by move=> i id; bw; apply nonempty_rep; apply hyp.
Qed.
Lemma product_nonempty2: forall f,
fgraph f -> nonempty (productb f) ->
(forall i, inc i (domain f) -> nonempty (V i f)).
Proof.
move=> f fgf [x]; aw; move=> [fg []] -> h i id; move:(h _ id) => w; ex_tac.
Qed.
Lemma productt_nonempty: forall (I:Set) (f:I->Set),
(forall i, nonempty (f i)) -> nonempty (productt f).
Proof.
rewrite /productt=> I f ane; apply product_nonempty.
apply gbcreate_fgraph.
move=> i; rewrite gbcreate_domain=> iI.
by rewrite - (B_eq iI) gbcreate_V; apply ane.
Qed.
Lemma productt_nonempty2: forall (I:Set) (f:I->Set),
nonempty (productt f) -> (forall i, nonempty (f i)).
Proof.
move=> I f [x xp] i; exists (V (Ro i) x); move: xp; aw; move=>[_[_]]; apply.
Qed.
fgraph f -> (forall i, inc i (domain f) -> nonempty (V i f)) ->
nonempty (productb f).
Proof.
move=> f fg hyp.
exists (L(domain f)(fun i=> rep (V i f))).
aw; split; gprops; split; bw.
by move=> i id; bw; apply nonempty_rep; apply hyp.
Qed.
Lemma product_nonempty2: forall f,
fgraph f -> nonempty (productb f) ->
(forall i, inc i (domain f) -> nonempty (V i f)).
Proof.
move=> f fgf [x]; aw; move=> [fg []] -> h i id; move:(h _ id) => w; ex_tac.
Qed.
Lemma productt_nonempty: forall (I:Set) (f:I->Set),
(forall i, nonempty (f i)) -> nonempty (productt f).
Proof.
rewrite /productt=> I f ane; apply product_nonempty.
apply gbcreate_fgraph.
move=> i; rewrite gbcreate_domain=> iI.
by rewrite - (B_eq iI) gbcreate_V; apply ane.
Qed.
Lemma productt_nonempty2: forall (I:Set) (f:I->Set),
nonempty (productt f) -> (forall i, nonempty (f i)).
Proof.
move=> I f [x xp] i; exists (V (Ro i) x); move: xp; aw; move=>[_[_]]; apply.
Qed.
The set of graphs is a product
Lemma graphset_pr1: forall a b z,
inc z (set_of_gfunctions a b) =
(fgraph z & domain z = a & sub z (product a b)).
Proof.
rewrite /set_of_gfunctions=> a b z.
apply iff_eq; rewrite Z_rw;aw; intuition.
Qed.
Lemma graphset_pr2: forall a b z,
inc z (set_of_gfunctions a b) =
(fgraph z & domain z = a & sub (range z) b).
Proof.
move=> a b z; rewrite graphset_pr1; apply iff_eq; rewrite corr_propcc.
intuition.
move=> [fgz [dz sr]]; rewrite dz; fprops.
Qed.
Lemma product_sub_graphset: forall f x,
fgraph f -> sub (unionb f) x ->
sub (productb f) (set_of_gfunctions (domain f) x).
Proof.
move=> f x fgf su t; aw; move=> [ft [dt hyp]]; rewrite graphset_pr2.
intuition.
move=> a; srw; move=> [b [bt]]->; apply su.
by apply unionb_inc with b; [ ue | eapply hyp ].
Qed.
Lemma product_eq_graphset: forall f x,
fgraph f -> (forall i, inc i (domain f) -> V i f = x) ->
productb f = set_of_gfunctions (domain f) x.
Proof.
move=> f x fgf prop; apply extensionality.
apply product_sub_graphset =>//.
by move=> y; srw; move=> [z [zd zV]]; rewrite - (prop _ zd).
move=> y; rewrite graphset_pr2; aw.
move=> [fy [dy sr]]; split=>//; split=>//.
move=> i id.
have r: inc (V i y) (range y) by fprops.
by rewrite dy in id; rewrite (prop _ id); apply sr.
Qed.
inc z (set_of_gfunctions a b) =
(fgraph z & domain z = a & sub z (product a b)).
Proof.
rewrite /set_of_gfunctions=> a b z.
apply iff_eq; rewrite Z_rw;aw; intuition.
Qed.
Lemma graphset_pr2: forall a b z,
inc z (set_of_gfunctions a b) =
(fgraph z & domain z = a & sub (range z) b).
Proof.
move=> a b z; rewrite graphset_pr1; apply iff_eq; rewrite corr_propcc.
intuition.
move=> [fgz [dz sr]]; rewrite dz; fprops.
Qed.
Lemma product_sub_graphset: forall f x,
fgraph f -> sub (unionb f) x ->
sub (productb f) (set_of_gfunctions (domain f) x).
Proof.
move=> f x fgf su t; aw; move=> [ft [dt hyp]]; rewrite graphset_pr2.
intuition.
move=> a; srw; move=> [b [bt]]->; apply su.
by apply unionb_inc with b; [ ue | eapply hyp ].
Qed.
Lemma product_eq_graphset: forall f x,
fgraph f -> (forall i, inc i (domain f) -> V i f = x) ->
productb f = set_of_gfunctions (domain f) x.
Proof.
move=> f x fgf prop; apply extensionality.
apply product_sub_graphset =>//.
by move=> y; srw; move=> [z [zd zV]]; rewrite - (prop _ zd).
move=> y; rewrite graphset_pr2; aw.
move=> [fy [dy sr]]; split=>//; split=>//.
move=> i id.
have r: inc (V i y) (range y) by fprops.
by rewrite dy in id; rewrite (prop _ id); apply sr.
Qed.
Product of a single set
Definition product1 (x a:Set) := productt (fun _:singleton a => x).
Definition cst_graph x y:= L x (fun _ => y).
Lemma product1_pr: forall x a,
product1 x a = set_of_gfunctions (singleton a) x.
Proof.
rewrite /product1 /productt=> x a.
move: (gbcreate_fgraph (fun _ : singleton a => x)) => f.
rewrite (product_eq_graphset f (x:=x)) gbcreate_domain //.
by move=> i ia; rewrite -(B_eq ia) gbcreate_V.
Qed.
Lemma product1_pr2: forall f x, fgraph f -> domain f = singleton x ->
product1 (V x f) x = productb f.
Proof.
move=> f x fgf df.
suff: (gbcreate (fun _ : singleton x => V x f) = f).
by rewrite /product1 /productt=> ->.
apply fgraph_exten=>//.
by apply gbcreate_fgraph.
by rewrite gbcreate_domain.
move=>y; rewrite gbcreate_domain=> xd.
rewrite -(B_eq xd) gbcreate_V (B_eq xd); awi xd; ue.
Qed.
Lemma product1_rw: forall x a y,
inc y (product1 x a) = (fgraph y & domain y = singleton a
& inc (V a y) x).
Proof.
move=> x a y; rewrite product1_pr graphset_pr2.
apply iff_eq; move=> [fgy [dy hyp]]; split=>//; split=>//.
apply hyp; apply inc_V_range=>//; rewrite dy; apply singleton_inc.
by move=> z; srw; rewrite dy; move=> [t []]; aw => -> ->.
Qed.
Definition product1_canon x a :=
BL (fun i : Set => cst_graph (singleton a) i) x (product1 x a).
Lemma product1_canon_axioms: forall x a,
transf_axioms (fun i => cst_graph (singleton a) i) x (product1 x a).
Proof.
move=> x a y yx; rewrite product1_rw /cst_graph; intuition; bw; fprops.
Qed.
Lemma product1_canon_function: forall x a,
is_function (product1_canon x a).
Proof.
rewrite /product1_canon=> x a; apply bl_function; apply product1_canon_axioms.
Qed.
Lemma product1_canon_W: forall x a i,
inc i x -> W i (product1_canon x a) = cst_graph (singleton a) i.
Proof.
by rewrite /product1_canon=> x a i iw; aw; apply product1_canon_axioms.
Qed.
Lemma cst_graph_pr: forall x y,
productb (cst_graph x y) = set_of_gfunctions x y.
Proof.
move=> x y.
have xd :x = domain (cst_graph x y) by rewrite /cst_graph; bw.
rewrite xd - product_eq_graphset.
by rewrite -xd.
rewrite /cst_graph; gprops.
rewrite -xd /cst_graph=> i ix; bw.
Qed.
Lemma product1_canon_bijective: forall x a,
bijection (product1_canon x a).
Proof.
rewrite /product1_canon=> x a; apply bl_bijective.
by apply product1_canon_axioms.
move=> u v ux vx eq; move: (f_equal (V a) eq); rewrite /cst_graph;bw; aw.
move=> y; rewrite product1_rw; move=> [fy [dy Vx]].
ex_tac.
rewrite -dy /cst_graph; apply fgraph_exten; gprops; bw.
by move=> z; rewrite dy=> zd; bw;move: zd; aw=>->.
Qed.
A product of n sets (with n=2) is isomorphic to the product of two sets
Definition product2 x y :=
productf two_points (variant TPa x y).
Definition product2_canon x y :=
BL (fun z => (variantLc (P z) (Q z))) (product x y) (product2 x y).
Lemma variant_if_rw1: forall x y, variant TPa x y TPa = x.
Proof. move=> x y; rewrite variant_if_rw //. Qed.
Lemma variant_if_not_rw1: forall x y, variant TPa x y TPb = y.
Proof.
move=> x y; rewrite variant_if_not_rw //; apply two_points_distinctb. Qed.
Hint Rewrite variant_if_not_rw1 variant_if_rw1 : bw.
Lemma product2_rw: forall x y z,
inc z (product2 x y) = (fgraph z & domain z = two_points &
inc (V TPa z) x &
inc (V TPb z) y).
Proof.
rewrite/product2=> x y z; aw.
apply iff_eq; move=> [fz [dz p]]; split=>//; split=>//.
rewrite dz in p.
move: (p _ inc_TPa_two_points); bw.
move: (p _ inc_TPb_two_points); bw.
by auto.
rewrite dz; move=> i ind; try_lvariant ind; intuition.
Qed.
Lemma product2_canon_axioms: forall x y,
transf_axioms (fun z => variantLc (P z) (Q z))
(product x y) (product2 x y).
Proof.
move=> x y t ; rewrite product2_rw; aw;move=> [pt [px qy]].
intuition; fprops;bw.
Qed.
Lemma product2_canon_function: forall x y,
is_function (product2_canon x y).
Proof.
rewrite/product2_canon=> x y.
apply bl_function; apply product2_canon_axioms.
Qed.
Lemma product2_canon_W: forall x y z,
inc z (product x y) -> W z (product2_canon x y) = variantLc (P z) (Q z).
Proof.
rewrite /product2_canon=> x y z zp; aw; awi zp=>//.
apply product2_canon_axioms.
Qed.
Lemma product2_canon_bijective: forall x y,
bijection (product2_canon x y).
Proof.
rewrite /product2_canon=> x y.
move: (product2_canon_axioms (x:=x) (y:=y)) => pa.
apply bl_bijective=>//.
move => u v; aw; move=> [up [pux quy]][vp [pvx qvy]] eq.
move: (f_equal (V TPa) eq)(f_equal (V TPb) eq); bw.
by apply pair_exten=>//.
move=> z; rewrite product2_rw; move => [fz [dz [Va Vb]]].
exists (J (V TPa z) (V TPb z)); split; first by aw; fprops.
apply fgraph_exten;fprops; bw; aw;try ue.
move=> t td; try_lvariant td.
Qed.
productf two_points (variant TPa x y).
Definition product2_canon x y :=
BL (fun z => (variantLc (P z) (Q z))) (product x y) (product2 x y).
Lemma variant_if_rw1: forall x y, variant TPa x y TPa = x.
Proof. move=> x y; rewrite variant_if_rw //. Qed.
Lemma variant_if_not_rw1: forall x y, variant TPa x y TPb = y.
Proof.
move=> x y; rewrite variant_if_not_rw //; apply two_points_distinctb. Qed.
Hint Rewrite variant_if_not_rw1 variant_if_rw1 : bw.
Lemma product2_rw: forall x y z,
inc z (product2 x y) = (fgraph z & domain z = two_points &
inc (V TPa z) x &
inc (V TPb z) y).
Proof.
rewrite/product2=> x y z; aw.
apply iff_eq; move=> [fz [dz p]]; split=>//; split=>//.
rewrite dz in p.
move: (p _ inc_TPa_two_points); bw.
move: (p _ inc_TPb_two_points); bw.
by auto.
rewrite dz; move=> i ind; try_lvariant ind; intuition.
Qed.
Lemma product2_canon_axioms: forall x y,
transf_axioms (fun z => variantLc (P z) (Q z))
(product x y) (product2 x y).
Proof.
move=> x y t ; rewrite product2_rw; aw;move=> [pt [px qy]].
intuition; fprops;bw.
Qed.
Lemma product2_canon_function: forall x y,
is_function (product2_canon x y).
Proof.
rewrite/product2_canon=> x y.
apply bl_function; apply product2_canon_axioms.
Qed.
Lemma product2_canon_W: forall x y z,
inc z (product x y) -> W z (product2_canon x y) = variantLc (P z) (Q z).
Proof.
rewrite /product2_canon=> x y z zp; aw; awi zp=>//.
apply product2_canon_axioms.
Qed.
Lemma product2_canon_bijective: forall x y,
bijection (product2_canon x y).
Proof.
rewrite /product2_canon=> x y.
move: (product2_canon_axioms (x:=x) (y:=y)) => pa.
apply bl_bijective=>//.
move => u v; aw; move=> [up [pux quy]][vp [pvx qvy]] eq.
move: (f_equal (V TPa) eq)(f_equal (V TPb) eq); bw.
by apply pair_exten=>//.
move=> z; rewrite product2_rw; move => [fz [dz [Va Vb]]].
exists (J (V TPa z) (V TPb z)); split; first by aw; fprops.
apply fgraph_exten;fprops; bw; aw;try ue.
move=> t td; try_lvariant td.
Qed.
The product of singletons is a singleton
Lemma product_singleton: forall f,
fgraph f -> (forall i, inc i (domain f) -> is_singleton (V i f))
-> is_singleton (productb f).
Proof.
move=> f ff als; rewrite is_singleton_rw; split.
apply product_nonempty=>//; move=> i id; move: (als _ id).
by rewrite is_singleton_rw; intuition.
move=> a b ap bp; apply (productb_exten ff ap bp)=> i id.
move: (als _ id);rewrite is_singleton_rw; move=> [ne eq].
move: ap bp; aw; move=> [_ [da ia]][_ [db ib]].
apply eq; [ apply ia | apply ib]; ue.
Qed.
Some properties of constant graphs
Definition is_constant_graph f :=
fgraph f &
(forall x x', inc x (domain f) -> inc x' (domain f) -> V x f = V x' f).
Lemma constant_graph_function: forall f,
is_constant_function f -> is_constant_graph(graph f).
Proof. move=> f [ff eq]; split; fprops; by move:ff=> [_[_]]<-. Qed.
Lemma constant_graph_V: forall s x y,
inc y s -> V y (cst_graph s x) = x.
Proof. rewrite /cst_graph=> s x y ys; bw. Qed.
Lemma constant_graph_small_range: forall f,
is_constant_graph f -> small_set(range f).
Proof.
by move=> f [ff eq] t1 t2; srw; move=> [x [xd]] -> [y [yd]] ->; apply eq.
Qed.
Definition diagonal_graphp e i :=
Zo(set_of_gfunctions i e) is_constant_graph.
Lemma diagonal_graph_rw: forall e i x,
inc x (diagonal_graphp e i) =
(is_constant_graph x & domain x = i & sub (range x) e).
Proof.
rewrite/diagonal_graphp /is_constant_graph=> e i x.
apply iff_eq; rewrite Z_rw graphset_pr2; intuition.
Qed.
Lemma constant_graph_is_constant: forall x y,
is_constant_graph (cst_graph x y).
Proof.
rewrite /is_constant_graph/cst_graph=> x y; aw.
split; [gprops | bw=> a b ad bd; bw].
Qed.
Definition constant_functor i e:=
BL(fun x => cst_graph i x) e (set_of_gfunctions i e).
Lemma cf_injective : forall i e,
nonempty i -> injection (constant_functor i e).
Proof.
rewrite /constant_functor;move=>i e [x xi].
apply bl_injective.
move=> t te; rewrite /cst_graph graphset_pr2.
split; first by gprops.
split; first by bw.
by move=> y; bw; move=> [b [be]] <-.
move=> u v ue ve eq; move: (f_equal (V x) eq).
by do 2 rewrite constant_graph_V //.
Qed.
fgraph f &
(forall x x', inc x (domain f) -> inc x' (domain f) -> V x f = V x' f).
Lemma constant_graph_function: forall f,
is_constant_function f -> is_constant_graph(graph f).
Proof. move=> f [ff eq]; split; fprops; by move:ff=> [_[_]]<-. Qed.
Lemma constant_graph_V: forall s x y,
inc y s -> V y (cst_graph s x) = x.
Proof. rewrite /cst_graph=> s x y ys; bw. Qed.
Lemma constant_graph_small_range: forall f,
is_constant_graph f -> small_set(range f).
Proof.
by move=> f [ff eq] t1 t2; srw; move=> [x [xd]] -> [y [yd]] ->; apply eq.
Qed.
Definition diagonal_graphp e i :=
Zo(set_of_gfunctions i e) is_constant_graph.
Lemma diagonal_graph_rw: forall e i x,
inc x (diagonal_graphp e i) =
(is_constant_graph x & domain x = i & sub (range x) e).
Proof.
rewrite/diagonal_graphp /is_constant_graph=> e i x.
apply iff_eq; rewrite Z_rw graphset_pr2; intuition.
Qed.
Lemma constant_graph_is_constant: forall x y,
is_constant_graph (cst_graph x y).
Proof.
rewrite /is_constant_graph/cst_graph=> x y; aw.
split; [gprops | bw=> a b ad bd; bw].
Qed.
Definition constant_functor i e:=
BL(fun x => cst_graph i x) e (set_of_gfunctions i e).
Lemma cf_injective : forall i e,
nonempty i -> injection (constant_functor i e).
Proof.
rewrite /constant_functor;move=>i e [x xi].
apply bl_injective.
move=> t te; rewrite /cst_graph graphset_pr2.
split; first by gprops.
split; first by bw.
by move=> y; bw; move=> [b [be]] <-.
move=> u v ue ve eq; move: (f_equal (V x) eq).
by do 2 rewrite constant_graph_V //.
Qed.
Change of variables in a product
Definition product_compose f u :=
BL (fun x => compose_graph x (graph u))
(productb f) (productf (source u) (fun k => V (W k u) f)).
Lemma compose_V: forall u v x,
fgraph u -> fgraph v -> fgraph (compose_graph u v) ->
inc x (domain (compose_graph u v)) ->
V x (compose_graph u v) = V (V x v) u.
Proof.
move=> u v x fu fgv fgc id.
move: (fdomain_pr1 fgc id).
by aw; move=> [_ [y [Jv Ju]]]; move:(pr2_V fgv Jv) (pr2_V fu Ju);aw=> <-.
Qed.
Lemma pc_axioms0: forall f u c,
fgraph f -> bijection u -> target u = domain f ->
inc c (productb f) ->
let g:= (corresp (domain c) (range c) c) \co u in
is_function g & compose_graph c (graph u) = graph g &
(forall i, inc i (source u) ->
V i (graph g) = V (V i (graph u)) c).
Proof.
move=> f u c ff bu tu; aw; move=> [fc [dc eq] g].
have fu: is_function u by fct_tac.
have cg: (compose_graph c (graph u) = graph g) by rewrite/g /compose; aw.
have fg: is_function g.
by rewrite/g; fct_tac; [ apply is_function_pr; fprops| aw; ue].
split=>//; split=>//.
move=> i iu; rewrite/g/compose; aw; rewrite compose_V //;
[fprops | ue | rewrite cg /g; fprops; aw].
Qed.
Lemma pc_axioms: forall f u,
fgraph f -> bijection u -> target u = domain f ->
transf_axioms (fun x => compose_graph x (graph u))
(productb f) (productf (source u) (fun k => V (W k u) f)).
Proof.
move=> f u ff bu tu x xp; move: (pc_axioms0 ff bu tu xp) => [fg [cg VV]].
have fu: is_function u by fct_tac.
aw; rewrite cg; split;first (by fprops); aw; split=>//.
move=> i iu.
rewrite (VV _ iu) -[ (V i (graph u))]/(W i u).
move:xp; aw; move=> [fx [dx]]; apply.
by rewrite dx -tu; apply inc_W_target.
Qed.
Lemma pc_function: forall f u,
fgraph f -> bijection u -> target u = domain f ->
is_function(product_compose f u).
Proof.
by rewrite/product_compose=> f iu ff bu tu; apply bl_function; apply pc_axioms.
Qed.
Lemma pc_W: forall f u x,
fgraph f -> bijection u -> target u = domain f ->
inc x (productb f) -> W x (product_compose f u) = compose_graph x (graph u).
Proof.
by rewrite /product_compose=> f u x ff bu tu xp; aw; apply pc_axioms.
Qed.
Lemma pc_WV: forall f u x i,
fgraph f -> bijection u -> target u = domain f ->
inc x (productb f) -> inc i (source u) ->
V i (W x (product_compose f u)) = V (W i u) x.
Proof.
move=> f u x i ff bu tu xp iu; rewrite pc_W //.
move:(pc_axioms0 ff bu tu xp) => [fg [cg VV]].
by rewrite cg VV.
Qed.
Lemma pc_bijective: forall f u,
fgraph f -> bijection u -> target u = domain f ->
bijection (product_compose f u).
Proof.
move=> f u ff bu tu.
move: (pc_function ff bu tu)=> fpc.
have su: surjection u by move: bu=> [_].
have spc: (source (product_compose f u) = productb f).
by rewrite /product_compose; aw.
split.
split=>//.
rewrite spc=> x y xs ys eq.
apply (productb_exten ff xs ys).
rewrite -tu; move=> i iu; move: (surjective_pr2 su iu) => [z [zs]] <-.
by rewrite - (pc_WV ff bu tu xs zs) -(pc_WV ff bu tu ys zs) eq.
apply surjective_pr6=>//.
set (nf:= L (source u) (fun k => V (W k u) f)).
have tpc:target (product_compose f u) = productb nf.
rewrite /product_compose; aw.
rewrite tpc spc => y yp.
set (x:= L (domain f) (fun i => V (W i (inverse_fun u)) y)).
have fnf: fgraph nf by rewrite /nf; gprops.
have xp: (inc x (productb f)).
rewrite /x; aw; bw; split; gprops; split=>//.
move=> i id; bw.
rewrite -tu in id; move: (surjective_pr2 su id) => [z [zsu wz]].
rewrite - (W_inverse2 bu zsu (sym_eq wz)) -wz.
move: yp; aw; move=> [fy [dy iVV]].
by rewrite dy /nf L_domain in iVV; move: (iVV _ zsu); bw.
ex_tac.
have Wt: inc (W x (product_compose f u)) (productb nf).
by rewrite - spc in xp; move: (inc_W_target fpc xp); rewrite tpc.
apply (productb_exten fnf Wt yp).
move=> i; rewrite /nf; bw=> id.
have Wd: inc (W i u)(domain f) by rewrite -tu; apply inc_W_target=>//; fct_tac.
rewrite pc_WV // /x; bw.
by rewrite -(W_inverse2 bu id (refl_equal (W i u))).
Qed.
BL (fun x => compose_graph x (graph u))
(productb f) (productf (source u) (fun k => V (W k u) f)).
Lemma compose_V: forall u v x,
fgraph u -> fgraph v -> fgraph (compose_graph u v) ->
inc x (domain (compose_graph u v)) ->
V x (compose_graph u v) = V (V x v) u.
Proof.
move=> u v x fu fgv fgc id.
move: (fdomain_pr1 fgc id).
by aw; move=> [_ [y [Jv Ju]]]; move:(pr2_V fgv Jv) (pr2_V fu Ju);aw=> <-.
Qed.
Lemma pc_axioms0: forall f u c,
fgraph f -> bijection u -> target u = domain f ->
inc c (productb f) ->
let g:= (corresp (domain c) (range c) c) \co u in
is_function g & compose_graph c (graph u) = graph g &
(forall i, inc i (source u) ->
V i (graph g) = V (V i (graph u)) c).
Proof.
move=> f u c ff bu tu; aw; move=> [fc [dc eq] g].
have fu: is_function u by fct_tac.
have cg: (compose_graph c (graph u) = graph g) by rewrite/g /compose; aw.
have fg: is_function g.
by rewrite/g; fct_tac; [ apply is_function_pr; fprops| aw; ue].
split=>//; split=>//.
move=> i iu; rewrite/g/compose; aw; rewrite compose_V //;
[fprops | ue | rewrite cg /g; fprops; aw].
Qed.
Lemma pc_axioms: forall f u,
fgraph f -> bijection u -> target u = domain f ->
transf_axioms (fun x => compose_graph x (graph u))
(productb f) (productf (source u) (fun k => V (W k u) f)).
Proof.
move=> f u ff bu tu x xp; move: (pc_axioms0 ff bu tu xp) => [fg [cg VV]].
have fu: is_function u by fct_tac.
aw; rewrite cg; split;first (by fprops); aw; split=>//.
move=> i iu.
rewrite (VV _ iu) -[ (V i (graph u))]/(W i u).
move:xp; aw; move=> [fx [dx]]; apply.
by rewrite dx -tu; apply inc_W_target.
Qed.
Lemma pc_function: forall f u,
fgraph f -> bijection u -> target u = domain f ->
is_function(product_compose f u).
Proof.
by rewrite/product_compose=> f iu ff bu tu; apply bl_function; apply pc_axioms.
Qed.
Lemma pc_W: forall f u x,
fgraph f -> bijection u -> target u = domain f ->
inc x (productb f) -> W x (product_compose f u) = compose_graph x (graph u).
Proof.
by rewrite /product_compose=> f u x ff bu tu xp; aw; apply pc_axioms.
Qed.
Lemma pc_WV: forall f u x i,
fgraph f -> bijection u -> target u = domain f ->
inc x (productb f) -> inc i (source u) ->
V i (W x (product_compose f u)) = V (W i u) x.
Proof.
move=> f u x i ff bu tu xp iu; rewrite pc_W //.
move:(pc_axioms0 ff bu tu xp) => [fg [cg VV]].
by rewrite cg VV.
Qed.
Lemma pc_bijective: forall f u,
fgraph f -> bijection u -> target u = domain f ->
bijection (product_compose f u).
Proof.
move=> f u ff bu tu.
move: (pc_function ff bu tu)=> fpc.
have su: surjection u by move: bu=> [_].
have spc: (source (product_compose f u) = productb f).
by rewrite /product_compose; aw.
split.
split=>//.
rewrite spc=> x y xs ys eq.
apply (productb_exten ff xs ys).
rewrite -tu; move=> i iu; move: (surjective_pr2 su iu) => [z [zs]] <-.
by rewrite - (pc_WV ff bu tu xs zs) -(pc_WV ff bu tu ys zs) eq.
apply surjective_pr6=>//.
set (nf:= L (source u) (fun k => V (W k u) f)).
have tpc:target (product_compose f u) = productb nf.
rewrite /product_compose; aw.
rewrite tpc spc => y yp.
set (x:= L (domain f) (fun i => V (W i (inverse_fun u)) y)).
have fnf: fgraph nf by rewrite /nf; gprops.
have xp: (inc x (productb f)).
rewrite /x; aw; bw; split; gprops; split=>//.
move=> i id; bw.
rewrite -tu in id; move: (surjective_pr2 su id) => [z [zsu wz]].
rewrite - (W_inverse2 bu zsu (sym_eq wz)) -wz.
move: yp; aw; move=> [fy [dy iVV]].
by rewrite dy /nf L_domain in iVV; move: (iVV _ zsu); bw.
ex_tac.
have Wt: inc (W x (product_compose f u)) (productb nf).
by rewrite - spc in xp; move: (inc_W_target fpc xp); rewrite tpc.
apply (productb_exten fnf Wt yp).
move=> i; rewrite /nf; bw=> id.
have Wd: inc (W i u)(domain f) by rewrite -tu; apply inc_W_target=>//; fct_tac.
rewrite pc_WV // /x; bw.
by rewrite -(W_inverse2 bu id (refl_equal (W i u))).
Qed.
a partial product is the set of restrictions
Lemma graph_exten: forall r r',
is_graph r -> is_graph r' ->
(r = r') = forall u v, related r u v = related r' u v.
Proof.
move=> r r' gr gr'; apply iff_eq; first by move => ->.
rewrite /related => p.
set_extens t => tr; [move: (gr _ tr)=> pt | move: (gr' _ tr)=> pt].
by rewrite -pt -p pt.
by rewrite -pt p pt.
Qed.
Lemma restriction_graph2 : forall f j,
fgraph f -> sub j (domain f) ->
L j (fun x => V x f) = compose_graph f (identity_g j).
Proof.
move=> f j fgj sj; rewrite restr_to_domain2 //.
rewrite graph_exten; fprops.
move=> u v; rewrite compose_related /restr/related Z_rw.
apply iff_eq.
by aw; move=> [pf uj]; exists u; aw.
by move=> [y []]; aw; move=> [uj] <-.
Qed.
is_graph r -> is_graph r' ->
(r = r') = forall u v, related r u v = related r' u v.
Proof.
move=> r r' gr gr'; apply iff_eq; first by move => ->.
rewrite /related => p.
set_extens t => tr; [move: (gr _ tr)=> pt | move: (gr' _ tr)=> pt].
by rewrite -pt -p pt.
by rewrite -pt p pt.
Qed.
Lemma restriction_graph2 : forall f j,
fgraph f -> sub j (domain f) ->
L j (fun x => V x f) = compose_graph f (identity_g j).
Proof.
move=> f j fgj sj; rewrite restr_to_domain2 //.
rewrite graph_exten; fprops.
move=> u v; rewrite compose_related /restr/related Z_rw.
apply iff_eq.
by aw; move=> [pf uj]; exists u; aw.
by move=> [y []]; aw; move=> [uj] <-.
Qed.
Product of f, with indices restricted to j
Definition restriction_product f j := productb (restr f j).
Definition pr_j f j :=
BL (fun z => restr z j) (productb f)(restriction_product f j).
Lemma prj_axioms: forall f j,
fgraph f -> sub j (domain f) ->
transf_axioms (fun z => restr z j)
(productb f)(restriction_product f j).
Proof.
rewrite /restriction_product=> f j ff sd x; aw; fprops; move=> [fx [dx iVV]].
have dr: j = domain (restr f j) by rewrite restr_domain1 //.
have sx: sub j (domain x) by rewrite dx.
rewrite -dr -restr_to_domain2 //; bw; split; gprops; split=>//.
by move=> i ij; bw; apply iVV; apply sx.
Qed.
Lemma prj_function: forall f j,
fgraph f -> sub j (domain f) -> is_function (pr_j f j).
Proof. by rewrite /pr_j => f j ff sd; apply bl_function; apply prj_axioms. Qed.
Lemma prj_W: forall f j x,
fgraph f -> sub j (domain f) -> inc x (productb f)
-> W x (pr_j f j) = (restr x j).
Proof. by rewrite /pr_j => f j x ff sd ix; aw; apply prj_axioms. Qed.
Lemma prj_WV: forall f j x i,
fgraph f -> sub j (domain f) -> inc x (productb f) -> inc i j
-> V i (W x (pr_j f j)) = V i x.
Proof.
move=> f j x u ff jd xp ij; rewrite prj_W //.
by move: xp; aw; move=> [fw [dd _]]; bw; ue.
Qed.
Lemma extension_partial_product: forall f j g,
fgraph f -> (forall i, inc i (domain f) -> nonempty (V i f)) ->
fgraph g -> domain g = j -> sub j (domain f) ->
(forall i, inc i j -> inc (V i g) (V i f)) ->
exists h, domain h = domain f &
fgraph h & (forall i, inc i (domain f) -> inc (V i h) (V i f)) &
(forall i, inc i j -> V i h = V i g).
Proof.
move=> f j g ff alne fg dg sjdf iVV.
set (h:= union2 g (unionf (complement (domain f) j)
(fun i => (singleton (J i (rep (V i f))))))).
have hg: is_graph h.
rewrite/h=>t; aw; case; first by move: fg=> [gg _]; apply gg.
by srw; move=> [y [_]]; aw => ->; fprops.
have dh: domain h = domain f.
set_extens t.
rewrite domain_rw // /h; move=> [y]; rewrite union2_rw.
case; first by move=> Jg; apply sjdf; rewrite -dg; dw; fprops; ex_tac.
srw; move=> [y' [yc Js]]; awi Js.
by rewrite (pr1_def Js); move: yc; srw; intuition.
move => tdf; case (inc_or_not t j).
rewrite -dg /h; dw; fprops; move => [y Jg].
by exists y; apply union2_first.
move=> ntj;dw; exists (rep (V t f));rewrite /h; apply union2_second.
by srw; exists t; ee; srw; ee.
have fgh: fgraph h.
split =>//.
move => x y; rewrite /h; aw; case=> xp; case=> yp sP.
by move: fg=> [_ ]; apply.
move: yp; srw; move=> [z []]; aw; srw; move=> [_ nzj] yJ.
by move: (inc_pr1_domain xp); rewrite dg sP yJ; aw=> zj; contradiction.
move: xp; srw; move=> [z []]; aw; srw; move=> [_ nzj] yJ.
by move: (inc_pr1_domain yp); rewrite dg -sP yJ; aw=> zj; contradiction.
move: xp yp; srw.
move=> [y1 [_]]; aw=> eq1; move=> [y2 [_]]; aw=> eq2.
have: y1 = y2 by move: sP; rewrite eq1 eq2; aw.
by rewrite eq1 eq2; move=>->.
have sV: forall i, inc i j -> V i h = V i g.
have sgh: (sub g h) by move=> t; rewrite /h; aw; intuition.
by move=> i; rewrite -dg=> ij; symmetry; apply (sub_graph_ev fgh sgh ij).
exists h.
split=>//; split=>//; split=>//.
move=> i id.
case (inc_or_not i j)=> hyp; first by rewrite sV //; apply iVV.
have ic: (inc i (complement (domain f) j)) by srw.
have Jh: (inc (J i (rep (V i f))) h).
by rewrite /h; apply union2_second; union_tac.
move: (pr2_V fgh Jh); aw; move=> <-.
by apply nonempty_rep; apply alne.
Qed.
Lemma prj_surjective: forall f j,
fgraph f -> (forall i, inc i (domain f) -> nonempty (V i f)) ->
sub j (domain f) -> surjection (pr_j f j).
Proof.
move=> f j ff alne sjd.
apply surjective_pr6; first by apply prj_function.
have fr: fgraph (restr f j) by fprops.
rewrite /pr_j=> y; rewrite {1} /restriction_product; aw.
rewrite - restr_to_domain2 //; bw; move=> [fy [dy iVVy]].
have iVV: (forall i, inc i j -> inc (V i y) (V i f)).
by move => i; rewrite -dy=> ij; move: (iVVy _ ij); bw; ue.
move: (extension_partial_product ff alne fy dy sjd iVV)
=> [h [dh [fgh [iVVh sv]]]].
have hp: (inc h (productb f)) by aw; rewrite dh; intuition.
ex_tac; rewrite prj_W //.
have sjdh: sub j (domain h) by rewrite dh.
symmetry; apply fgraph_exten; fprops.
rewrite dy restr_domain1 //.
by rewrite dy; move=> x xd; bw; symmetry; apply sv.
Qed.
Lemma pri_surjective: forall f k,
fgraph f -> (forall i, inc i (domain f) -> nonempty (V i f)) ->
inc k (domain f) -> surjection (pr_i f k).
Proof.
move=> f k ff alne kd.
set (j:= singleton k).
have sjd: (sub j (domain f)).
by move=> t; rewrite/j singleton_rw => ->.
have cpc: (product1_canon (V k f) k) \coP (pr_i f k).
split; first by apply product1_canon_function.
split; first by apply pri_function.
by rewrite /product1_canon /pr_i; aw.
have fcpx:(is_function ((product1_canon (V k f) k) \co (pr_i f k))).
by fct_tac.
have prjc: (pr_j f j = (product1_canon (V k f) k) \co (pr_i f k)).
apply function_exten =>//.
by apply prj_function=>//.
by rewrite /pr_j/pr_i; aw.
rewrite /pr_j/pr_i/product1_canon/product1; aw.
have: restr f j = cst_graph (singleton k) (V k f).
rewrite/j /cst_graph.
apply fgraph_exten;fprops;gprops;bw;rewrite restr_domain1 //.
by move=> x eq; bw; rewrite (singleton_eq eq).
by rewrite/restriction_product; move=>->.
move=> x xs.
have xp: inc x (productb f) by move: xs;rewrite /pr_j bl_source.
have xsk: inc x (source (pr_i f k)) by rewrite /pr_i bl_source.
have aux1: inc (W x (pr_i f k)) (V k f).
have : (target (pr_i f k) = V k f) by rewrite/pr_i; aw.
by move=> <-; apply inc_W_target=>//; apply pri_function=>//.
aw=>//; rewrite product1_canon_W //.
rewrite prj_W //.
have xp1: inc x (productb f) by [].
move: xp1; aw;move => [fx [dx iVVx]].
have ssk: sub (singleton k) (domain x).
by move=> xw eq1; bw; rewrite dx (singleton_eq eq1).
rewrite/j; apply fgraph_exten;
rewrite /cst_graph ;fprops;gprops;bw; rewrite restr_domain1 //.
by move=> z zk; bw; rewrite (singleton_eq zk) pri_W //.
apply (surj_left_compose2 cpc).
by rewrite - prjc; apply prj_surjective.
move: (product1_canon_bijective (V k f) k)=> [ h _]; apply h.
Qed.
Lemma productb_monotone1: forall f g,
fgraph f -> fgraph g -> domain f = domain g ->
(forall i, inc i (domain f) -> sub (V i f) (V i g))
-> sub (productb f) (productb g).
Proof.
move=> f g ff fg dfdg sVV x; aw; move=> [fx [dx iVV]].
by intuition; [ ue| apply sVV ; [ue| apply iVV]].
Qed.
Lemma productb_monotone2: forall f g,
fgraph f -> fgraph g -> domain f = domain g ->
(forall i, inc i (domain f) -> nonempty (V i f)) ->
sub (productb f) (productb g) ->
(forall i, inc i (domain f) -> sub (V i f) (V i g)).
Proof.
move=> f g ff fg dfdg alne sfg i id j jV.
have jt: (inc j (target (pr_i f i))) by rewrite/ pr_i; aw.
move:(surjective_pr2 (pri_surjective ff alne id) jt) => [x [xs]] <-.
move: xs; rewrite /pr_i bl_source=>xs.
rewrite pri_W //.
by move: (sfg _ xs); aw; move=> [fx [dx]]; apply; rewrite dx -dfdg.
Qed.
Definition pr_j f j :=
BL (fun z => restr z j) (productb f)(restriction_product f j).
Lemma prj_axioms: forall f j,
fgraph f -> sub j (domain f) ->
transf_axioms (fun z => restr z j)
(productb f)(restriction_product f j).
Proof.
rewrite /restriction_product=> f j ff sd x; aw; fprops; move=> [fx [dx iVV]].
have dr: j = domain (restr f j) by rewrite restr_domain1 //.
have sx: sub j (domain x) by rewrite dx.
rewrite -dr -restr_to_domain2 //; bw; split; gprops; split=>//.
by move=> i ij; bw; apply iVV; apply sx.
Qed.
Lemma prj_function: forall f j,
fgraph f -> sub j (domain f) -> is_function (pr_j f j).
Proof. by rewrite /pr_j => f j ff sd; apply bl_function; apply prj_axioms. Qed.
Lemma prj_W: forall f j x,
fgraph f -> sub j (domain f) -> inc x (productb f)
-> W x (pr_j f j) = (restr x j).
Proof. by rewrite /pr_j => f j x ff sd ix; aw; apply prj_axioms. Qed.
Lemma prj_WV: forall f j x i,
fgraph f -> sub j (domain f) -> inc x (productb f) -> inc i j
-> V i (W x (pr_j f j)) = V i x.
Proof.
move=> f j x u ff jd xp ij; rewrite prj_W //.
by move: xp; aw; move=> [fw [dd _]]; bw; ue.
Qed.
Lemma extension_partial_product: forall f j g,
fgraph f -> (forall i, inc i (domain f) -> nonempty (V i f)) ->
fgraph g -> domain g = j -> sub j (domain f) ->
(forall i, inc i j -> inc (V i g) (V i f)) ->
exists h, domain h = domain f &
fgraph h & (forall i, inc i (domain f) -> inc (V i h) (V i f)) &
(forall i, inc i j -> V i h = V i g).
Proof.
move=> f j g ff alne fg dg sjdf iVV.
set (h:= union2 g (unionf (complement (domain f) j)
(fun i => (singleton (J i (rep (V i f))))))).
have hg: is_graph h.
rewrite/h=>t; aw; case; first by move: fg=> [gg _]; apply gg.
by srw; move=> [y [_]]; aw => ->; fprops.
have dh: domain h = domain f.
set_extens t.
rewrite domain_rw // /h; move=> [y]; rewrite union2_rw.
case; first by move=> Jg; apply sjdf; rewrite -dg; dw; fprops; ex_tac.
srw; move=> [y' [yc Js]]; awi Js.
by rewrite (pr1_def Js); move: yc; srw; intuition.
move => tdf; case (inc_or_not t j).
rewrite -dg /h; dw; fprops; move => [y Jg].
by exists y; apply union2_first.
move=> ntj;dw; exists (rep (V t f));rewrite /h; apply union2_second.
by srw; exists t; ee; srw; ee.
have fgh: fgraph h.
split =>//.
move => x y; rewrite /h; aw; case=> xp; case=> yp sP.
by move: fg=> [_ ]; apply.
move: yp; srw; move=> [z []]; aw; srw; move=> [_ nzj] yJ.
by move: (inc_pr1_domain xp); rewrite dg sP yJ; aw=> zj; contradiction.
move: xp; srw; move=> [z []]; aw; srw; move=> [_ nzj] yJ.
by move: (inc_pr1_domain yp); rewrite dg -sP yJ; aw=> zj; contradiction.
move: xp yp; srw.
move=> [y1 [_]]; aw=> eq1; move=> [y2 [_]]; aw=> eq2.
have: y1 = y2 by move: sP; rewrite eq1 eq2; aw.
by rewrite eq1 eq2; move=>->.
have sV: forall i, inc i j -> V i h = V i g.
have sgh: (sub g h) by move=> t; rewrite /h; aw; intuition.
by move=> i; rewrite -dg=> ij; symmetry; apply (sub_graph_ev fgh sgh ij).
exists h.
split=>//; split=>//; split=>//.
move=> i id.
case (inc_or_not i j)=> hyp; first by rewrite sV //; apply iVV.
have ic: (inc i (complement (domain f) j)) by srw.
have Jh: (inc (J i (rep (V i f))) h).
by rewrite /h; apply union2_second; union_tac.
move: (pr2_V fgh Jh); aw; move=> <-.
by apply nonempty_rep; apply alne.
Qed.
Lemma prj_surjective: forall f j,
fgraph f -> (forall i, inc i (domain f) -> nonempty (V i f)) ->
sub j (domain f) -> surjection (pr_j f j).
Proof.
move=> f j ff alne sjd.
apply surjective_pr6; first by apply prj_function.
have fr: fgraph (restr f j) by fprops.
rewrite /pr_j=> y; rewrite {1} /restriction_product; aw.
rewrite - restr_to_domain2 //; bw; move=> [fy [dy iVVy]].
have iVV: (forall i, inc i j -> inc (V i y) (V i f)).
by move => i; rewrite -dy=> ij; move: (iVVy _ ij); bw; ue.
move: (extension_partial_product ff alne fy dy sjd iVV)
=> [h [dh [fgh [iVVh sv]]]].
have hp: (inc h (productb f)) by aw; rewrite dh; intuition.
ex_tac; rewrite prj_W //.
have sjdh: sub j (domain h) by rewrite dh.
symmetry; apply fgraph_exten; fprops.
rewrite dy restr_domain1 //.
by rewrite dy; move=> x xd; bw; symmetry; apply sv.
Qed.
Lemma pri_surjective: forall f k,
fgraph f -> (forall i, inc i (domain f) -> nonempty (V i f)) ->
inc k (domain f) -> surjection (pr_i f k).
Proof.
move=> f k ff alne kd.
set (j:= singleton k).
have sjd: (sub j (domain f)).
by move=> t; rewrite/j singleton_rw => ->.
have cpc: (product1_canon (V k f) k) \coP (pr_i f k).
split; first by apply product1_canon_function.
split; first by apply pri_function.
by rewrite /product1_canon /pr_i; aw.
have fcpx:(is_function ((product1_canon (V k f) k) \co (pr_i f k))).
by fct_tac.
have prjc: (pr_j f j = (product1_canon (V k f) k) \co (pr_i f k)).
apply function_exten =>//.
by apply prj_function=>//.
by rewrite /pr_j/pr_i; aw.
rewrite /pr_j/pr_i/product1_canon/product1; aw.
have: restr f j = cst_graph (singleton k) (V k f).
rewrite/j /cst_graph.
apply fgraph_exten;fprops;gprops;bw;rewrite restr_domain1 //.
by move=> x eq; bw; rewrite (singleton_eq eq).
by rewrite/restriction_product; move=>->.
move=> x xs.
have xp: inc x (productb f) by move: xs;rewrite /pr_j bl_source.
have xsk: inc x (source (pr_i f k)) by rewrite /pr_i bl_source.
have aux1: inc (W x (pr_i f k)) (V k f).
have : (target (pr_i f k) = V k f) by rewrite/pr_i; aw.
by move=> <-; apply inc_W_target=>//; apply pri_function=>//.
aw=>//; rewrite product1_canon_W //.
rewrite prj_W //.
have xp1: inc x (productb f) by [].
move: xp1; aw;move => [fx [dx iVVx]].
have ssk: sub (singleton k) (domain x).
by move=> xw eq1; bw; rewrite dx (singleton_eq eq1).
rewrite/j; apply fgraph_exten;
rewrite /cst_graph ;fprops;gprops;bw; rewrite restr_domain1 //.
by move=> z zk; bw; rewrite (singleton_eq zk) pri_W //.
apply (surj_left_compose2 cpc).
by rewrite - prjc; apply prj_surjective.
move: (product1_canon_bijective (V k f) k)=> [ h _]; apply h.
Qed.
Lemma productb_monotone1: forall f g,
fgraph f -> fgraph g -> domain f = domain g ->
(forall i, inc i (domain f) -> sub (V i f) (V i g))
-> sub (productb f) (productb g).
Proof.
move=> f g ff fg dfdg sVV x; aw; move=> [fx [dx iVV]].
by intuition; [ ue| apply sVV ; [ue| apply iVV]].
Qed.
Lemma productb_monotone2: forall f g,
fgraph f -> fgraph g -> domain f = domain g ->
(forall i, inc i (domain f) -> nonempty (V i f)) ->
sub (productb f) (productb g) ->
(forall i, inc i (domain f) -> sub (V i f) (V i g)).
Proof.
move=> f g ff fg dfdg alne sfg i id j jV.
have jt: (inc j (target (pr_i f i))) by rewrite/ pr_i; aw.
move:(surjective_pr2 (pri_surjective ff alne id) jt) => [x [xs]] <-.
move: xs; rewrite /pr_i bl_source=>xs.
rewrite pri_W //.
by move: (sfg _ xs); aw; move=> [fx [dx]]; apply; rewrite dx -dfdg.
Qed.
Given a product with domain I, and a partition J of I, we can consider the
set of mappings that associate to each j in J the partial product.
This is a product. Associativity says that this product of products
is isomorphic to the initial product.
Implies associativity of cardinal product
Definition prod_assoc_axioms f g :=
fgraph f & partition_fam g (domain f).
Definition prod_assoc_map f g :=
BL(fun z => (L (domain g) (fun l => W z (pr_j f (V l g)))))
(productb f)
(productf (domain g) (fun l => (restriction_product f (V l g)))).
Lemma pam_axioms: forall f g,
prod_assoc_axioms f g ->
transf_axioms(fun z => (L (domain g) (fun l => W z (pr_j f (V l g)))))
(productb f)
(productf (domain g) (fun l => (restriction_product f (V l g)))).
Proof.
move=> f g [fgf [fg [md u]]] t td; aw; split; bw; gprops; split=>//.
move=> i id; bw.
have: (restriction_product f (V i g) = target (pr_j f (V i g))).
by rewrite /pr_j; aw.
move=>->; apply inc_W_target; last by rewrite/pr_j bl_source.
apply prj_function=>//.
move=> z; rewrite -u=> tu; union_tac.
Qed.
Lemma pam_function: forall f g,
prod_assoc_axioms f g ->
is_function (prod_assoc_map f g).
Proof.
by rewrite /prod_assoc_map=>f g a; apply bl_function; apply pam_axioms.
Qed.
Lemma pam_W: forall f g x,
prod_assoc_axioms f g -> inc x (productb f) ->
W x (prod_assoc_map f g) = (L (domain g) (fun l => W x (pr_j f (V l g)))).
Proof. by rewrite /prod_assoc_map=> f g x a xp; aw; apply pam_axioms. Qed.
Lemma pam_injective: forall f g,
prod_assoc_axioms f g ->
injection (prod_assoc_map f g).
Proof.
move=> f g pa; split; first by apply pam_function.
have sp:(source (prod_assoc_map f g) = productb f).
by rewrite /prod_assoc_map; aw.
rewrite sp {sp} => x y xs ys.
(do 2 rewrite pam_W //)=> sv.
move: pa=> [ff [fg [md ugdf]]].
apply (@productb_exten _ x y ff)=>//.
move=> i; rewrite - ugdf; srw; move=> [z [zd iV]].
have svf: sub (V z g) (domain f) by move=> t tv; rewrite - ugdf; union_tac.
move: (f_equal (V z) sv); bw; rewrite prj_W // prj_W // => sv1.
move: xs ys (f_equal (V i) sv1); aw; move=> [fx [dx _]] [fy [dy _]]; bw; ue.
Qed.
Theorem pam_bijective: forall f g,
prod_assoc_axioms f g ->
bijection (prod_assoc_map f g).
Proof.
move=> f g pa; split; first by apply pam_injective.
apply surjective_pr6; first by apply pam_function.
have sp: source (prod_assoc_map f g) = productb f
by rewrite /prod_assoc_map; aw.
have tp: target (prod_assoc_map f g) =
(productf (domain g) (fun l : Set => restriction_product f (V l g))).
by rewrite /prod_assoc_map; aw.
rewrite sp tp; move=> y; move=> yp; move: (yp);aw; move=> [fy [dy iVr]].
move: (pa) => [fgf pfgd].
move: (pfgd)=> [fg [md ugdf]].
have sVd: forall i, inc i (domain g) -> sub (V i g) (domain f).
by move=> i id v vV; rewrite -ugdf; union_tac.
pose h i := BL (fun z => (V z (V i y))) (V i g) (unionb f).
have ta:forall i, inc i (domain g) ->
transf_axioms (fun z => V z (V i y)) (V i g) (unionb f).
move=> i id x xV; move: (sVd _ id) =>sVdi.
have tmp: fgraph (restr f (V i g)) by fprops.
apply unionb_inc with x; first by apply sVdi=>//.
rewrite -dy in id; move: (iVr _ id).
rewrite /restriction_product; aw.
rewrite - restr_to_domain2 //; bw.
by move=> [fgV [deq rel]]; rewrite deq in rel; move: (rel _ xV); bw.
have afp: forall i, inc i (domain g) -> function_prop (h i) (V i g) (unionb f).
move=> i id; rewrite /h.
by split; first (by apply bl_function; apply ta); aw.
move: (extension_partition pfgd afp) => [[x [[fx [sx tx]] ag]] _].
have gxp: inc (graph x) (productb f).
aw;split; fprops; split =>//.
move=> i ix; rewrite -[V i (graph x)]/(W i x).
move: ix; rewrite sx -ugdf; srw; move=>[z [zdg iV]].
move: (ag _ zdg)=> [_ [_ eq]]; rewrite (eq _ iV) /h.
aw; last by apply (ta _ zdg).
have aux: fgraph (restr f (V z g)) by fprops.
have aux2: sub (V z g) (domain f) by move=> t tV; rewrite -ugdf; union_tac.
rewrite -dy in zdg; move: (iVr _ zdg).
rewrite /restriction_product; aw.
rewrite - restr_to_domain2 //; bw.
by move=> [_ [eq1 res]]; rewrite eq1 in res; move: (res _ iV); bw.
have Wt: inc (W (graph x) (prod_assoc_map f g))(target (prod_assoc_map f g)).
by apply inc_W_target; [ by apply pam_function | ue ].
ex_tac.
symmetry.
apply (productf_exten (x:=y) (x':= W _ _ ) yp).
by rewrite - tp.
move=> i id.
have aux2: sub (V i g) (domain f) by apply sVd.
have aux3: sub (V i g) (domain (graph x)) by aw; ue.
rewrite pam_W=> //; bw; rewrite prj_W //.
rewrite - dy in id; move: (iVr _ id); rewrite/restriction_product; aw; fprops.
move=> [fgy [dy2 rel]].
rewrite restr_domain1 in dy2=>//.
apply fgraph_exten; fprops.
rewrite dy2 restr_domain1=>//; fprops.
move=> z zd; rewrite dy2 in zd; bw; fprops.
rewrite -[V z (graph x)]/(W z x).
rewrite dy in id; move: (ag _ id).
move=> [ aa [bb cc]]; rewrite (cc _ zd).
by rewrite /h; aw; apply ta.
Qed.
Lemma variantLc_prop: forall x y,
variantLc x y = L two_points (variant TPa x y).
Proof. by move=> x y; rewrite/ variantLc/variantL two_points_pr2. Qed.
Ltac eqtrans u:= apply equipotent_transitive with u; fprops.
Ltac eqsym:= apply equipotent_symmetric.
Lemma prod_assoc_map2: forall f g,
prod_assoc_axioms f g -> domain g = two_points
-> (productb f) \Eq
(product (restriction_product f (V TPa g))(restriction_product f (V TPb g))).
Proof.
move=> f g pam dg.
eqtrans (productf (domain g) (fun l => (restriction_product f (V l g)))).
exists (prod_assoc_map f g).
split; first by apply pam_bijective.
by rewrite /prod_assoc_map; aw.
eqsym.
exists (product2_canon (restriction_product f (V TPa g))
(restriction_product f (V TPb g))).
split; first by apply product2_canon_bijective.
rewrite /product2_canon; aw; split=>//.
rewrite /product2/productf.
apply f_equal; apply fgraph_exten;gprops; bw.
by symmetry.
rewrite dg; move=> x xt; try_lvariant xt; fprops.
Qed.
fgraph f & partition_fam g (domain f).
Definition prod_assoc_map f g :=
BL(fun z => (L (domain g) (fun l => W z (pr_j f (V l g)))))
(productb f)
(productf (domain g) (fun l => (restriction_product f (V l g)))).
Lemma pam_axioms: forall f g,
prod_assoc_axioms f g ->
transf_axioms(fun z => (L (domain g) (fun l => W z (pr_j f (V l g)))))
(productb f)
(productf (domain g) (fun l => (restriction_product f (V l g)))).
Proof.
move=> f g [fgf [fg [md u]]] t td; aw; split; bw; gprops; split=>//.
move=> i id; bw.
have: (restriction_product f (V i g) = target (pr_j f (V i g))).
by rewrite /pr_j; aw.
move=>->; apply inc_W_target; last by rewrite/pr_j bl_source.
apply prj_function=>//.
move=> z; rewrite -u=> tu; union_tac.
Qed.
Lemma pam_function: forall f g,
prod_assoc_axioms f g ->
is_function (prod_assoc_map f g).
Proof.
by rewrite /prod_assoc_map=>f g a; apply bl_function; apply pam_axioms.
Qed.
Lemma pam_W: forall f g x,
prod_assoc_axioms f g -> inc x (productb f) ->
W x (prod_assoc_map f g) = (L (domain g) (fun l => W x (pr_j f (V l g)))).
Proof. by rewrite /prod_assoc_map=> f g x a xp; aw; apply pam_axioms. Qed.
Lemma pam_injective: forall f g,
prod_assoc_axioms f g ->
injection (prod_assoc_map f g).
Proof.
move=> f g pa; split; first by apply pam_function.
have sp:(source (prod_assoc_map f g) = productb f).
by rewrite /prod_assoc_map; aw.
rewrite sp {sp} => x y xs ys.
(do 2 rewrite pam_W //)=> sv.
move: pa=> [ff [fg [md ugdf]]].
apply (@productb_exten _ x y ff)=>//.
move=> i; rewrite - ugdf; srw; move=> [z [zd iV]].
have svf: sub (V z g) (domain f) by move=> t tv; rewrite - ugdf; union_tac.
move: (f_equal (V z) sv); bw; rewrite prj_W // prj_W // => sv1.
move: xs ys (f_equal (V i) sv1); aw; move=> [fx [dx _]] [fy [dy _]]; bw; ue.
Qed.
Theorem pam_bijective: forall f g,
prod_assoc_axioms f g ->
bijection (prod_assoc_map f g).
Proof.
move=> f g pa; split; first by apply pam_injective.
apply surjective_pr6; first by apply pam_function.
have sp: source (prod_assoc_map f g) = productb f
by rewrite /prod_assoc_map; aw.
have tp: target (prod_assoc_map f g) =
(productf (domain g) (fun l : Set => restriction_product f (V l g))).
by rewrite /prod_assoc_map; aw.
rewrite sp tp; move=> y; move=> yp; move: (yp);aw; move=> [fy [dy iVr]].
move: (pa) => [fgf pfgd].
move: (pfgd)=> [fg [md ugdf]].
have sVd: forall i, inc i (domain g) -> sub (V i g) (domain f).
by move=> i id v vV; rewrite -ugdf; union_tac.
pose h i := BL (fun z => (V z (V i y))) (V i g) (unionb f).
have ta:forall i, inc i (domain g) ->
transf_axioms (fun z => V z (V i y)) (V i g) (unionb f).
move=> i id x xV; move: (sVd _ id) =>sVdi.
have tmp: fgraph (restr f (V i g)) by fprops.
apply unionb_inc with x; first by apply sVdi=>//.
rewrite -dy in id; move: (iVr _ id).
rewrite /restriction_product; aw.
rewrite - restr_to_domain2 //; bw.
by move=> [fgV [deq rel]]; rewrite deq in rel; move: (rel _ xV); bw.
have afp: forall i, inc i (domain g) -> function_prop (h i) (V i g) (unionb f).
move=> i id; rewrite /h.
by split; first (by apply bl_function; apply ta); aw.
move: (extension_partition pfgd afp) => [[x [[fx [sx tx]] ag]] _].
have gxp: inc (graph x) (productb f).
aw;split; fprops; split =>//.
move=> i ix; rewrite -[V i (graph x)]/(W i x).
move: ix; rewrite sx -ugdf; srw; move=>[z [zdg iV]].
move: (ag _ zdg)=> [_ [_ eq]]; rewrite (eq _ iV) /h.
aw; last by apply (ta _ zdg).
have aux: fgraph (restr f (V z g)) by fprops.
have aux2: sub (V z g) (domain f) by move=> t tV; rewrite -ugdf; union_tac.
rewrite -dy in zdg; move: (iVr _ zdg).
rewrite /restriction_product; aw.
rewrite - restr_to_domain2 //; bw.
by move=> [_ [eq1 res]]; rewrite eq1 in res; move: (res _ iV); bw.
have Wt: inc (W (graph x) (prod_assoc_map f g))(target (prod_assoc_map f g)).
by apply inc_W_target; [ by apply pam_function | ue ].
ex_tac.
symmetry.
apply (productf_exten (x:=y) (x':= W _ _ ) yp).
by rewrite - tp.
move=> i id.
have aux2: sub (V i g) (domain f) by apply sVd.
have aux3: sub (V i g) (domain (graph x)) by aw; ue.
rewrite pam_W=> //; bw; rewrite prj_W //.
rewrite - dy in id; move: (iVr _ id); rewrite/restriction_product; aw; fprops.
move=> [fgy [dy2 rel]].
rewrite restr_domain1 in dy2=>//.
apply fgraph_exten; fprops.
rewrite dy2 restr_domain1=>//; fprops.
move=> z zd; rewrite dy2 in zd; bw; fprops.
rewrite -[V z (graph x)]/(W z x).
rewrite dy in id; move: (ag _ id).
move=> [ aa [bb cc]]; rewrite (cc _ zd).
by rewrite /h; aw; apply ta.
Qed.
Lemma variantLc_prop: forall x y,
variantLc x y = L two_points (variant TPa x y).
Proof. by move=> x y; rewrite/ variantLc/variantL two_points_pr2. Qed.
Ltac eqtrans u:= apply equipotent_transitive with u; fprops.
Ltac eqsym:= apply equipotent_symmetric.
Lemma prod_assoc_map2: forall f g,
prod_assoc_axioms f g -> domain g = two_points
-> (productb f) \Eq
(product (restriction_product f (V TPa g))(restriction_product f (V TPb g))).
Proof.
move=> f g pam dg.
eqtrans (productf (domain g) (fun l => (restriction_product f (V l g)))).
exists (prod_assoc_map f g).
split; first by apply pam_bijective.
by rewrite /prod_assoc_map; aw.
eqsym.
exists (product2_canon (restriction_product f (V TPa g))
(restriction_product f (V TPb g))).
split; first by apply product2_canon_bijective.
rewrite /product2_canon; aw; split=>//.
rewrite /product2/productf.
apply f_equal; apply fgraph_exten;gprops; bw.
by symmetry.
rewrite dg; move=> x xt; try_lvariant xt; fprops.
Qed.
The product of two sets is equipotent to one of them if the other one is a singleton. Implies that one is a unit for carinal product.
Lemma first_proj_bijective: forall x y,
is_singleton y -> bijection (first_proj (product x y)).
Proof.
move=> x y sy.
set f:=first_proj (product x y).
have gp: is_graph (product x y) by apply product_is_graph.
have sp: source f = (product x y) by rewrite /f/first_proj; aw.
have tp: target f = domain(product x y) by rewrite /f/first_proj; aw.
have fp: is_function (first_proj (product x y)) by apply first_proj_function.
move: sy=> [t ty].
split.
split=>//; rewrite sp=> a b ap bp.
do 2 rewrite first_proj_W=>//.
move: ap bp; rewrite ty; aw.
by move=> [ap [_ qa]] [bp [_ qb]] peq; apply pair_exten=>//; ue.
apply surjective_pr6=> //.
rewrite sp tp; move=>z; rewrite domain_rw//;move=> [w wp].
ex_tac; rewrite first_proj_W=>//; aw.
Qed.
A product is isomophic to a partial product if missing factors are singletons. Implies that one can be removed in cardinal products.
Lemma prj_bijective: forall f j,
fgraph f -> sub j (domain f) ->
(forall i, inc i (complement (domain f) j) -> is_singleton (V i f)) ->
bijection (pr_j f j).
Proof.
move=> f j fgf sjd als.
set (k:= (complement (domain f) j)) in *.
have sk: sub k (domain f) by apply sub_complement.
have dk: domain (restr f k) = k.
by rewrite - restr_to_domain2=>//; bw.
have sr: is_singleton (restriction_product f k).
rewrite /restriction_product; apply product_singleton; fprops.
by rewrite dk => i ik; bw; apply als.
set (g:= variantLc j (complement (domain f) j)).
have paa: prod_assoc_axioms f g.
split=>//; apply (is_partition_with_complement sjd).
set (f1:= prod_assoc_map f g).
have bf1: bijection f1 by apply pam_bijective.
have ff1:is_function f1 by fct_tac.
set (xa:= restriction_product f (V TPa g)).
set (xb:= restriction_product f (V TPb g)).
set (f2:= inverse_fun (product2_canon xa xb)).
have dg: domain g = two_points by rewrite /g; bw.
have Va: V TPa g = j by rewrite /g; bw.
have Vb: V TPb g = k by rewrite /g; bw.
have bf2: bijection f2.
by apply inverse_bij_is_bij; apply product2_canon_bijective.
have ff2: is_function f2 by fct_tac.
have sf2tf1: source f2 = target f1 .
rewrite /f1/f2/product2_canon/prod_assoc_map; aw.
rewrite/product2 dg /productf.
by apply f_equal; apply L_exten1=>//; move=> v vt; try_lvariant vt.
set (f3:= first_proj (product xa xb)).
have sb: is_singleton xb.
by rewrite /xb /restriction_product Vb; apply sr.
have bf3: bijection f3 by apply first_proj_bijective.
have sf3tf2: source f3 = target f2.
by rewrite /f3/f2/product2_canon /first_proj; aw.
have cf3f3: f3 \coP f2 by rewrite /f3/f2; split=>//; fct_tac.
have bcf3f2: bijection (f3 \co f2) by apply compose_bijective.
have cf3f2f1: (f3 \co f2) \coP f1 by split; aw=>//; fct_tac.
suff: (pr_j f j = (f3 \co f2) \co f1).
by move=> ->; apply compose_bijective=>//.
have sp: source (pr_j f j) = source f1.
by rewrite /pr_j/f1/prod_assoc_map; aw.
have nexb: nonempty xb by move: sb; rewrite is_singleton_rw; intuition.
apply function_exten; aw.
by apply prj_function=>//.
by fct_tac.
by rewrite/pr_j/f3/first_proj product_domain=>//; aw; rewrite /xa Va.
rewrite sp=> x xs.
move xW: (W x f1)=>y.
move yW: (W y f2)=>z.
have yt: inc y (target f1) by rewrite -xW; fprops.
rewrite -sf2tf1 in yt.
have zt: inc z (target f2) by rewrite -yW; apply inc_W_target=>//.
have Jg2: (inc (J y z) (graph f2)) by rewrite -yW; graph_tac.
have fp: is_function (product2_canon xa xb) by apply product2_canon_function.
have zs: inc z (source (product2_canon xa xb)).
by move: zt; rewrite /f2 /inverse_fun; aw.
have zp: inc z (product xa xb) by move: zs; rewrite /product2_canon; aw.
aw; last by rewrite xW.
rewrite xW yW first_proj_W //.
move: Jg2; rewrite /f2/inverse_fun; aw=> Jg3.
move: (W_pr2 fp Jg3); aw; rewrite product2_canon_W=>//.
move=> eq; move: (f_equal (V TPa) eq); bw; move => <-.
rewrite -xW /f1 pam_W //; bw; try rewrite Va =>//.
by rewrite dg; fprops.
by move: xs; rewrite/f1/prod_assoc_map bl_source.
Qed.
Given a double family Xij, union over I of intersection over J is
intersection of union, and conversely. The new index set is a product.
Theorem distrib_union_inter: forall f,
fgraph f -> nonempty (domain f) ->
(forall l, inc l (domain f) -> fgraph (V l f)) ->
(forall l, inc l (domain f) -> nonempty (domain (V l f))) ->
unionf (domain f) (fun l => intersectionb (V l f)) =
intersectionf (productf (domain f) (fun l => (domain (V l f))))
(fun g => (unionf (domain f) (fun l => V (V l g) (V l f)))).
Proof.
move=> f ff ned afg alne.
set_extens s => xu.
apply intersectionf_inc.
rewrite/productf; apply product_nonempty; gprops.
by bw;move=> i id; bw; apply alne.
move: (unionf_exists xu) => [y [yf si]].
move => j; aw; move=> [fgj [dj aiVd]].
union_tac.
by rewrite dj in aiVd; apply (intersectionb_forall si (aiVd _ yf)).
ex_middle nsu.
have ne: (nonempty (productf (domain f) (fun l =>
(Zo (domain (V l f)) (fun i => ~ (inc s (V i (V l f)))))))).
rewrite /productf; apply product_nonempty; gprops; bw.
move=> i id; bw; set jl:=Zo _ _ .
case (emptyset_dichot jl) =>//; move=> je; elim nsu; union_tac.
apply intersectionb_inc; first by rewrite -nonempty_domain; apply alne.
move=> j jd; ex_middle hyp.
by empty_tac1 j; rewrite/jl Z_rw.
move: ne=> [y]. aw. move => [fy [dy iVz]].
have yp: inc y (productf (domain f) (fun l => domain (V l f))).
aw; split =>//; split =>// i id.
by move: (iVz _ id); rewrite Z_rw; intuition.
move: (intersectionf_forall xu yp).
srw; move=> [z [zdf sv]]; rewrite - dy in zdf; move: (iVz _ zdf).
rewrite Z_rw; move=> [_ nsv] ; contradiction.
Qed.
Theorem distrib_inter_union: forall f,
fgraph f -> nonempty (domain f) ->
(forall l, inc l (domain f) -> fgraph (V l f)) ->
(forall l, inc l (domain f) -> nonempty (domain (V l f))) ->
intersectionf (domain f) (fun l => unionb (V l f)) =
unionf (productf (domain f) (fun l => (domain (V l f))))
(fun g => (intersectionf (domain f) (fun l => V (V l g) (V l f)))).
Proof.
move=> f ff ne alfg alne; set_extens x => xi; last first.
apply intersectionf_inc=>//; move => j jd.
move: (unionf_exists xi) => [y []]; aw; move => [_ [dy Vd]] xi1.
by move: (intersectionf_forall xi1 jd)=> u; union_tac; apply Vd; ue.
pose yf j:= choose (fun a => inc a (domain (V j f))& inc x (V a (V j f))).
have p1: (forall j, inc j (domain f) -> (inc (yf j) (domain (V j f))&
inc x (V (yf j) (V j f)))).
by move=> j jd; apply choose_pr; move:(intersectionf_forall xi jd); srw.
apply unionf_inc with (L (domain f) yf).
aw; split; gprops; bw; split=>//.
by bw; move=> i id; bw; move: (p1 _ id); intuition.
by apply intersectionf_inc=>//; move=> j jd; bw; move: (p1 _ jd); intuition.
Qed.
fgraph f -> nonempty (domain f) ->
(forall l, inc l (domain f) -> fgraph (V l f)) ->
(forall l, inc l (domain f) -> nonempty (domain (V l f))) ->
unionf (domain f) (fun l => intersectionb (V l f)) =
intersectionf (productf (domain f) (fun l => (domain (V l f))))
(fun g => (unionf (domain f) (fun l => V (V l g) (V l f)))).
Proof.
move=> f ff ned afg alne.
set_extens s => xu.
apply intersectionf_inc.
rewrite/productf; apply product_nonempty; gprops.
by bw;move=> i id; bw; apply alne.
move: (unionf_exists xu) => [y [yf si]].
move => j; aw; move=> [fgj [dj aiVd]].
union_tac.
by rewrite dj in aiVd; apply (intersectionb_forall si (aiVd _ yf)).
ex_middle nsu.
have ne: (nonempty (productf (domain f) (fun l =>
(Zo (domain (V l f)) (fun i => ~ (inc s (V i (V l f)))))))).
rewrite /productf; apply product_nonempty; gprops; bw.
move=> i id; bw; set jl:=Zo _ _ .
case (emptyset_dichot jl) =>//; move=> je; elim nsu; union_tac.
apply intersectionb_inc; first by rewrite -nonempty_domain; apply alne.
move=> j jd; ex_middle hyp.
by empty_tac1 j; rewrite/jl Z_rw.
move: ne=> [y]. aw. move => [fy [dy iVz]].
have yp: inc y (productf (domain f) (fun l => domain (V l f))).
aw; split =>//; split =>// i id.
by move: (iVz _ id); rewrite Z_rw; intuition.
move: (intersectionf_forall xu yp).
srw; move=> [z [zdf sv]]; rewrite - dy in zdf; move: (iVz _ zdf).
rewrite Z_rw; move=> [_ nsv] ; contradiction.
Qed.
Theorem distrib_inter_union: forall f,
fgraph f -> nonempty (domain f) ->
(forall l, inc l (domain f) -> fgraph (V l f)) ->
(forall l, inc l (domain f) -> nonempty (domain (V l f))) ->
intersectionf (domain f) (fun l => unionb (V l f)) =
unionf (productf (domain f) (fun l => (domain (V l f))))
(fun g => (intersectionf (domain f) (fun l => V (V l g) (V l f)))).
Proof.
move=> f ff ne alfg alne; set_extens x => xi; last first.
apply intersectionf_inc=>//; move => j jd.
move: (unionf_exists xi) => [y []]; aw; move => [_ [dy Vd]] xi1.
by move: (intersectionf_forall xi1 jd)=> u; union_tac; apply Vd; ue.
pose yf j:= choose (fun a => inc a (domain (V j f))& inc x (V a (V j f))).
have p1: (forall j, inc j (domain f) -> (inc (yf j) (domain (V j f))&
inc x (V (yf j) (V j f)))).
by move=> j jd; apply choose_pr; move:(intersectionf_forall xi jd); srw.
apply unionf_inc with (L (domain f) yf).
aw; split; gprops; bw; split=>//.
by bw; move=> i id; bw; move: (p1 _ id); intuition.
by apply intersectionf_inc=>//; move=> j jd; bw; move: (p1 _ jd); intuition.
Qed.
Variants of distributivity of union and intersection of Xij
when one index set has two elements
Lemma distrib_union2_inter: forall f g,
fgraph f -> fgraph g -> nonempty (domain f) -> nonempty (domain g) ->
union2 (intersectionb f)(intersectionb g) =
intersectionf(product (domain f)(domain g)) (fun z =>
(union2 (V (P z) f) (V (Q z) g))).
Proof.
move=> f g ff fg nef neg.
have nep:nonempty (product (domain f) (domain g)).
by move: nef neg=> [x xd][y yd]; exists (J x y); fprops.
set_extens x => xu.
apply intersectionf_inc =>//.
move=> j; rewrite product_inc_rw; move=> [pj [p1 p2]].
case (union2_or xu)=>p.
by move: (intersectionb_forall p p1)=>a; apply union2_first.
by move: (intersectionb_forall p p2)=>a; apply union2_second.
case (inc_or_not x (intersectionb f))=> p; first by apply union2_first.
have: (exists z1, inc z1 (domain f) & ~ (inc x (V z1 f))).
apply exists_proof; dneg np.
apply intersectionb_inc; first by rewrite -nonempty_domain.
move=> i id; case (inc_or_not x (V i f))=>// => neq; elim (np i); auto.
move =>[z1 [z1d nexv]].
apply union2_second.
apply intersectionb_inc;first by rewrite -nonempty_domain.
move=>i id.
move: (intersectionf_forall xu (product_pair_inc z1d id)); aw; case=>//.
Qed.
Lemma distrib_inter2_union: forall f g,
fgraph f -> fgraph g -> nonempty (domain f) -> nonempty (domain g) ->
intersection2 (unionb f)(unionb g) =
unionf(product (domain f)(domain g)) (fun z =>
(intersection2 (V (P z) f) (V (Q z) g))).
Proof.
move=> f g ff fg nedf nedg.
set_extens x.
aw; srw; move => [[y [yd xf]] [z [zd xg]]].
by exists (J y z); split; fprops; aw.
srw; move=> [y]; rewrite product_inc_rw; move=> [[yp [yf yg]]].
aw; move=> [p1 p2]; split; union_tac.
Qed.
Some small lemmas
Lemma trivial_product1: forall f, productf emptyset f = singleton emptyset.
Proof.
rewrite/productf=> f; rewrite trivial_fgraph; apply product_trivial.
Qed.
Lemma unionf_emptyset: forall f, unionf emptyset f = emptyset.
Proof. by move=> f; apply empty_unionf. Qed.
Lemma nonempty_product3: forall sf f i,
inc i sf -> f i = emptyset -> productf sf f = emptyset.
Proof.
move=> sf f i.
case (emptyset_dichot (productf sf f)) =>//.
move=> [x]; aw; move=> [fx []] -> p isf fi.
move: (p _ isf); rewrite fi => bad; elim (emptyset_pr bad).
Qed.
Distributivity of product over union and intersection
Theorem distrib_prod_union: forall f,
fgraph f ->
(forall l, inc l (domain f) -> fgraph (V l f)) ->
productf (domain f) (fun l => unionb (V l f)) =
unionf (productf (domain f) (fun l => (domain (V l f))))
(fun g => (productf (domain f) (fun l => V (V l g) (V l f)))).
Proof.
move=> f ff alfg.
case (emptyset_dichot (domain f)) => p.
by rewrite p 2! trivial_product1 unionf_singleton trivial_product1.
case (p_or_not_p (forall l, inc l (domain f) -> nonempty (domain (V l f)))).
move=> alned.
set_extens x => xp.
have : (nonempty (productf (domain f) (fun l =>
(Zo (domain (V l f)) (fun i => (inc (V l x) (V i (V l f)))))))).
apply product_nonempty; gprops; bw.
move=> i id; bw.
move: xp; aw; move=> [gx []] -> px; move: (alfg _ id) (px _ id)=> fgx.
by srw; move=> [y [yd iVV]]; exists y; apply Z_inc.
move=> [y]; move: xp; aw; move=> [fx [dx px]] [fy [dy py]].
apply unionf_inc with y; aw; split=>//; split=>//.
by move=> i id; move: (py _ id); rewrite Z_rw; intuition.
by move=> i; rewrite dx -dy=> id;move: (py _ id); rewrite Z_rw; intuition.
move: xp;srw; move=> [y]; aw; move => [[fy [dy py]] [fx [dx px]]].
split=>//; split=>//; move=> i id.
by apply unionb_inc with (V i y); [ apply py; rewrite dy -dx | apply px].
move=> special.
have : exists l, inc l (domain f) & domain (V l f) = emptyset.
apply exists_proof; dneg aux; move=> l ld.
case (emptyset_dichot (domain (V l f)))=>//; move=>pp.
by elim (aux l); intuition.
move=> [x [xd de]].
rewrite(@nonempty_product3 _ (fun l => domain (V l f)) _ xd de).
rewrite unionf_emptyset.
have ue: (unionb (V x f) = emptyset).
by rewrite/unionb de; apply unionf_emptyset.
by rewrite(@nonempty_product3 _ (fun l => unionb (V l f)) _ xd ue).
Qed.
Theorem distrib_prod_intersection: forall f,
fgraph f -> nonempty (domain f) ->
(forall l, inc l (domain f) -> fgraph (V l f)) ->
(forall l, inc l (domain f) -> nonempty (domain (V l f))) ->
productf (domain f) (fun l => intersectionb (V l f)) =
intersectionf (productf (domain f) (fun l => (domain (V l f))))
(fun g => (productf (domain f) (fun l => V (V l g) (V l f)))).
Proof.
move=> f ff ne alf alne.
have nep:(nonempty (productb (L (domain f) (fun l => domain (V l f))))).
rewrite/ productf; apply product_nonempty; gprops; bw.
by move=> i id; bw; apply alne.
set_extens x => xp.
apply intersectionf_inc=>//.
move => j; move: xp; aw; move=>[fx [xd iVi]] [fj [ dj iVd]].
split=>//;split=>//; move => i id; move: (iVi _ id) => ivdi.
by apply intersectionb_forall=>//; apply iVd; rewrite dj -xd.
move: nep=> [u up]. move: (intersectionf_forall xp up).
aw; move=> [fx [dx iVx]]; split =>//; split=>//.
move=> i id.
apply intersectionb_inc.
by rewrite -nonempty_domain; apply alne; ue.
move=> j jd.
have: (nonempty (productf (domain f) (fun l =>
(Zo (domain (V l f)) (fun jj => l = i -> jj = j))))).
rewrite /productf; apply product_nonempty; gprops; bw.
move=> k kd; bw.
case (equal_or_not k i)=> ki.
by exists j; apply Z_inc; ue.
move: (alne _ kd) => [y yd]; exists y; apply Z_inc=>//.
move=> [y]; aw; move=> [fy [dy iV]].
have yd: (inc y (productf (domain f) (fun l => domain (V l f)))).
aw; split=>//; split=>//.
by move=> k kd; move: (iV _ kd); rewrite Z_rw; move => [iVd kip] =>//.
move: (intersectionf_forall xp yd); aw; move => [_ [_ rel]].
rewrite dy -dx in iV; move: (iV _ id); rewrite Z_rw.
by move=> [_] <- //; apply rel.
Qed.
Lemma partition_product: forall f,
fgraph f -> nonempty (domain f) ->
(forall l, inc l (domain f) -> fgraph (V l f)) ->
(forall l, inc l (domain f) -> (partition_fam (V l f) (unionb (V l f)))) ->
partition_fam(L(productf (domain f) (fun l => domain (V l f)) )
(fun g => (productf (domain f) (fun l => V (V l g) (V l f)))))
( productf (domain f) (fun l => unionb (V l f))).
Proof.
move=> f ff ned alg alp.
split; first (by gprops); split.
move=> i j; rewrite L_domain => ip jp; bw.
eq_dichot ij; red; set int:= intersection2 _ _.
case: (emptyset_dichot int) =>// ei.
elim ij; move: ei=> [xi]; rewrite/int; aw.
move: ip jp; aw.
move=> [fi [di vi]] [fj [dj vj]] [[fx [dx eq1]] [_ [_ eq2]]].
apply fgraph_exten =>//; try ue.
move=> x xd; rewrite dx in eq1 eq2; rewrite di in xd.
move: (eq1 _ xd) (eq2 _ xd) (alp _ xd).
rewrite /partition_fam; move=> eq3 eq4 [fgV [md _]].
rewrite di in vi; move: (vi _ xd)=> eq5.
rewrite dj in vj; move: (vj _ xd)=> eq6.
move: (md _ _ eq5 eq6).
case=>//.
rewrite/disjoint=> dij; empty_tac1 (V x xi); by apply intersection2_inc.
rewrite distrib_prod_union /unionb=>//.
bw; apply unionf_exten =>//.
by move => i ip; bw.
Qed.
fgraph f ->
(forall l, inc l (domain f) -> fgraph (V l f)) ->
productf (domain f) (fun l => unionb (V l f)) =
unionf (productf (domain f) (fun l => (domain (V l f))))
(fun g => (productf (domain f) (fun l => V (V l g) (V l f)))).
Proof.
move=> f ff alfg.
case (emptyset_dichot (domain f)) => p.
by rewrite p 2! trivial_product1 unionf_singleton trivial_product1.
case (p_or_not_p (forall l, inc l (domain f) -> nonempty (domain (V l f)))).
move=> alned.
set_extens x => xp.
have : (nonempty (productf (domain f) (fun l =>
(Zo (domain (V l f)) (fun i => (inc (V l x) (V i (V l f)))))))).
apply product_nonempty; gprops; bw.
move=> i id; bw.
move: xp; aw; move=> [gx []] -> px; move: (alfg _ id) (px _ id)=> fgx.
by srw; move=> [y [yd iVV]]; exists y; apply Z_inc.
move=> [y]; move: xp; aw; move=> [fx [dx px]] [fy [dy py]].
apply unionf_inc with y; aw; split=>//; split=>//.
by move=> i id; move: (py _ id); rewrite Z_rw; intuition.
by move=> i; rewrite dx -dy=> id;move: (py _ id); rewrite Z_rw; intuition.
move: xp;srw; move=> [y]; aw; move => [[fy [dy py]] [fx [dx px]]].
split=>//; split=>//; move=> i id.
by apply unionb_inc with (V i y); [ apply py; rewrite dy -dx | apply px].
move=> special.
have : exists l, inc l (domain f) & domain (V l f) = emptyset.
apply exists_proof; dneg aux; move=> l ld.
case (emptyset_dichot (domain (V l f)))=>//; move=>pp.
by elim (aux l); intuition.
move=> [x [xd de]].
rewrite(@nonempty_product3 _ (fun l => domain (V l f)) _ xd de).
rewrite unionf_emptyset.
have ue: (unionb (V x f) = emptyset).
by rewrite/unionb de; apply unionf_emptyset.
by rewrite(@nonempty_product3 _ (fun l => unionb (V l f)) _ xd ue).
Qed.
Theorem distrib_prod_intersection: forall f,
fgraph f -> nonempty (domain f) ->
(forall l, inc l (domain f) -> fgraph (V l f)) ->
(forall l, inc l (domain f) -> nonempty (domain (V l f))) ->
productf (domain f) (fun l => intersectionb (V l f)) =
intersectionf (productf (domain f) (fun l => (domain (V l f))))
(fun g => (productf (domain f) (fun l => V (V l g) (V l f)))).
Proof.
move=> f ff ne alf alne.
have nep:(nonempty (productb (L (domain f) (fun l => domain (V l f))))).
rewrite/ productf; apply product_nonempty; gprops; bw.
by move=> i id; bw; apply alne.
set_extens x => xp.
apply intersectionf_inc=>//.
move => j; move: xp; aw; move=>[fx [xd iVi]] [fj [ dj iVd]].
split=>//;split=>//; move => i id; move: (iVi _ id) => ivdi.
by apply intersectionb_forall=>//; apply iVd; rewrite dj -xd.
move: nep=> [u up]. move: (intersectionf_forall xp up).
aw; move=> [fx [dx iVx]]; split =>//; split=>//.
move=> i id.
apply intersectionb_inc.
by rewrite -nonempty_domain; apply alne; ue.
move=> j jd.
have: (nonempty (productf (domain f) (fun l =>
(Zo (domain (V l f)) (fun jj => l = i -> jj = j))))).
rewrite /productf; apply product_nonempty; gprops; bw.
move=> k kd; bw.
case (equal_or_not k i)=> ki.
by exists j; apply Z_inc; ue.
move: (alne _ kd) => [y yd]; exists y; apply Z_inc=>//.
move=> [y]; aw; move=> [fy [dy iV]].
have yd: (inc y (productf (domain f) (fun l => domain (V l f)))).
aw; split=>//; split=>//.
by move=> k kd; move: (iV _ kd); rewrite Z_rw; move => [iVd kip] =>//.
move: (intersectionf_forall xp yd); aw; move => [_ [_ rel]].
rewrite dy -dx in iV; move: (iV _ id); rewrite Z_rw.
by move=> [_] <- //; apply rel.
Qed.
Lemma partition_product: forall f,
fgraph f -> nonempty (domain f) ->
(forall l, inc l (domain f) -> fgraph (V l f)) ->
(forall l, inc l (domain f) -> (partition_fam (V l f) (unionb (V l f)))) ->
partition_fam(L(productf (domain f) (fun l => domain (V l f)) )
(fun g => (productf (domain f) (fun l => V (V l g) (V l f)))))
( productf (domain f) (fun l => unionb (V l f))).
Proof.
move=> f ff ned alg alp.
split; first (by gprops); split.
move=> i j; rewrite L_domain => ip jp; bw.
eq_dichot ij; red; set int:= intersection2 _ _.
case: (emptyset_dichot int) =>// ei.
elim ij; move: ei=> [xi]; rewrite/int; aw.
move: ip jp; aw.
move=> [fi [di vi]] [fj [dj vj]] [[fx [dx eq1]] [_ [_ eq2]]].
apply fgraph_exten =>//; try ue.
move=> x xd; rewrite dx in eq1 eq2; rewrite di in xd.
move: (eq1 _ xd) (eq2 _ xd) (alp _ xd).
rewrite /partition_fam; move=> eq3 eq4 [fgV [md _]].
rewrite di in vi; move: (vi _ xd)=> eq5.
rewrite dj in vj; move: (vj _ xd)=> eq6.
move: (md _ _ eq5 eq6).
case=>//.
rewrite/disjoint=> dij; empty_tac1 (V x xi); by apply intersection2_inc.
rewrite distrib_prod_union /unionb=>//.
bw; apply unionf_exten =>//.
by move => i ip; bw.
Qed.
Special cases when one index set has only two elements
Lemma distrib_prod2_union: forall f g,
fgraph f -> fgraph g ->
nonempty (domain f) -> nonempty (domain g) ->
product2 (unionb f)(unionb g) =
unionf(product (domain f)(domain g))
(fun z => (product2 (V (P z) f) (V (Q z) g))).
Proof.
move=> f g ff fg nef neg; rewrite /product2; set_extens x.
aw; move=> [fx [dx iV]] ; rewrite dx in iV.
have Tat: inc TPa two_points by fprops.
have Tbt: inc TPb two_points by fprops.
move: (iV _ Tat) (iV _ Tbt); bw; srw.
move=> [u [ud uVV]] [v [vd vVV]].
exists (J u v); split; fprops; aw; split=>//;split=>//.
by move=> i; rewrite dx => ind; try_lvariant ind.
srw; move=>[y []]; rewrite product_inc_rw; move=> [yp [py qy]].
aw; move =>[fx [dx iV]]; split=>//; split=>//.
move=> i idx; move: (iV _ idx); rewrite dx in idx.
try_lvariant idx; move=> ?; union_tac.
Qed.
Lemma distrib_prod2_inter: forall f g,
fgraph f -> fgraph g ->
nonempty (domain f) -> nonempty (domain g) ->
product2 (intersectionb f)(intersectionb g)=
intersectionf(product (domain f)(domain g)) (fun z =>
(product2 (V (P z) f) (V (Q z) g))).
Proof.
move=> f g ff fg nf ng.
have nep: nonempty (product (domain f) (domain g)).
by move: nf ng=> [x xf] [y yg]; exists (J x y); fprops.
rewrite/product2; set_extens x.
aw; move=> [fx [dx]]; rewrite dx; move=> iVv.
apply intersectionf_inc => //.
move => j; rewrite product_inc_rw; move => [jp [pj qj]]; aw.
split=>//; split=>//; move=> i ; rewrite dx => id; move: (iVv _ id).
try_lvariant id=> Vi.
by apply (intersectionf_forall Vi pj).
by apply (intersectionf_forall Vi qj).
rewrite intersectionf_rw=>// =>ip.
move: nep=> [j jp]; move: (ip _ jp)=> xp.
have aux: (forall i j, inc i (domain f) -> inc j (domain g) ->
(inc (V TPa x) (V i f) & inc (V TPb x) (V j g))).
move=> i k id kd.
have Jd: (inc (J i k) (product (domain f) (domain g))) by fprops.
move: (ip _ Jd); aw; move=> [fx [dx px]].
have ta: (inc TPa (domain x)) by ue.
have tb: (inc TPb (domain x)) by ue.
by move: (px _ ta)(px _ tb); bw; intuition.
move: xp; aw; move=> [fx [dx px]]; split=>//; split=>//.
move=> i; rewrite dx=> id.
try_lvariant id; apply intersectionb_inc; try rewrite -nonempty_domain //.
by move => k kd; move: ng=> [l lg]; move: (aux _ _ kd lg); intuition.
by move => l lg; move:nf=> [k kd]; move: (aux _ _ kd lg); intuition.
Qed.
Lemma distrib_product2_union: forall f g,
fgraph f -> fgraph g ->
product (unionb f)(unionb g) =
unionf(product (domain f)(domain g)) (fun z =>
(product (V (P z) f) (V (Q z) g))).
Proof.
move=> f g ff fg.
case (emptyset_dichot (domain f)) =>nf.
rewrite nf empty_unionf; srw.
by rewrite/unionb nf unionf_emptyset; srw.
case (emptyset_dichot (domain g)) => ng.
rewrite ng empty_unionf; srw.
by rewrite/unionb ng unionf_emptyset; srw.
move: (distrib_prod2_union ff fg nf ng); rewrite/product2 => distr.
set_extens x => xp.
move: xp;aw; move=> [xp [px qx]].
set (v := variantLc (P x) (Q x)).
have: inc v (productb (variantLc (unionb f) (unionb g))).
aw; fprops; rewrite/v; split; fprops; bw; split=>//.
by move=> i it; try_lvariant it.
rewrite /productf in distr; rewrite variantLc_prop distr.
srw; move=> [y [p1 p2]]; exists y; split=>//.
move: p2; aw; gprops; move=> [fv []]; bw => dv p.
have ta: (inc TPa (domain v)) by rewrite /v; bw; fprops.
have tb: (inc TPb (domain v)) by rewrite /v; bw; fprops.
by move: (p _ ta) (p _ tb); rewrite /v; bw; fprops.
move: xp; srw.
move=> [y [yp]]; rewrite 2! product_inc_rw; move=> [px [xp xq]].
set (qx := variantLc (P x) (Q x)).
have:(inc qx (productb (variantLc (unionb f) (unionb g)))).
move: distr; rewrite/productf -variantLc_prop; move=>->.
union_tac; aw; gprops;rewrite/ qx.
split; fprops; bw;split=>//.
move=> i it; try_lvariant it; fprops.
aw; fprops; rewrite variantLc_domain.
move=> [fq [dx vqx]].
have ta: (inc TPa (domain qx)) by rewrite dx; fprops.
have tb: (inc TPb (domain qx)) by rewrite dx; fprops.
move: (vqx _ ta) (vqx _ tb); rewrite /qx; bw.
intuition.
Qed.
Lemma distrib_product2_inter: forall f g,
fgraph f -> fgraph g -> nonempty (domain f) -> nonempty (domain g) ->
product (intersectionb f)(intersectionb g) =
intersectionf(product (domain f)(domain g)) (fun z =>
(product (V (P z) f) (V (Q z) g))).
Proof.
move=>f g ff fg nef neg.
move: (distrib_prod2_inter ff fg nef neg); rewrite / product2=> distr.
have nep: nonempty (product (domain f) (domain g)).
by move: nef neg; move=> [y yf][z zg]; exists (J y z); fprops.
set_extens x => xp.
set (qx := variantLc (P x) (Q x)).
have: inc qx (productb (variantLc (intersectionb f) (intersectionb g))).
move: xp; aw; gprops; fprops; move=>[xp [pxf qxg]].
rewrite /qx; split;fprops;bw; split=>//.
by move=> i ind; try_lvariant ind.
move: distr; rewrite /productf -variantLc_prop; move=> -> qxp.
apply intersectionf_inc=>//.
move=> j jp; move: (intersectionf_forall qxp jp).
move: xp; aw; gprops; move=> [px [xp xq]] [fgq [dq vq]].
move: dq; bw => dq; rewrite dq in vq.
have ta: (inc TPa two_points) by fprops.
have tb: (inc TPb two_points) by fprops.
move: (vq _ ta)(vq _ tb); rewrite /qx; bw.
by intuition.
set (qx := variantLc (P x) (Q x)).
have: inc qx (productb (variantLc (intersectionb f) (intersectionb g))).
move: distr; rewrite /productf variantLc_prop; move =>->.
apply intersectionf_inc=>//.
move=> j jp; move: (intersectionf_forall xp jp).
aw;gprops; move=> [_ [Px Qx]].
rewrite /qx; split;fprops; bw; split=>//.
move=> i it; try_lvariant it; fprops.
aw; fprops;move=> [_ [_]];rewrite /qx; bw => qv.
move: nep=> [w wp]; move: (intersectionf_forall xp wp); aw; move=> [px _].
split=>//.
move: (qv _ inc_TPa_two_points) (qv _ inc_TPb_two_points); bw.
intuition.
Qed.
fgraph f -> fgraph g ->
nonempty (domain f) -> nonempty (domain g) ->
product2 (unionb f)(unionb g) =
unionf(product (domain f)(domain g))
(fun z => (product2 (V (P z) f) (V (Q z) g))).
Proof.
move=> f g ff fg nef neg; rewrite /product2; set_extens x.
aw; move=> [fx [dx iV]] ; rewrite dx in iV.
have Tat: inc TPa two_points by fprops.
have Tbt: inc TPb two_points by fprops.
move: (iV _ Tat) (iV _ Tbt); bw; srw.
move=> [u [ud uVV]] [v [vd vVV]].
exists (J u v); split; fprops; aw; split=>//;split=>//.
by move=> i; rewrite dx => ind; try_lvariant ind.
srw; move=>[y []]; rewrite product_inc_rw; move=> [yp [py qy]].
aw; move =>[fx [dx iV]]; split=>//; split=>//.
move=> i idx; move: (iV _ idx); rewrite dx in idx.
try_lvariant idx; move=> ?; union_tac.
Qed.
Lemma distrib_prod2_inter: forall f g,
fgraph f -> fgraph g ->
nonempty (domain f) -> nonempty (domain g) ->
product2 (intersectionb f)(intersectionb g)=
intersectionf(product (domain f)(domain g)) (fun z =>
(product2 (V (P z) f) (V (Q z) g))).
Proof.
move=> f g ff fg nf ng.
have nep: nonempty (product (domain f) (domain g)).
by move: nf ng=> [x xf] [y yg]; exists (J x y); fprops.
rewrite/product2; set_extens x.
aw; move=> [fx [dx]]; rewrite dx; move=> iVv.
apply intersectionf_inc => //.
move => j; rewrite product_inc_rw; move => [jp [pj qj]]; aw.
split=>//; split=>//; move=> i ; rewrite dx => id; move: (iVv _ id).
try_lvariant id=> Vi.
by apply (intersectionf_forall Vi pj).
by apply (intersectionf_forall Vi qj).
rewrite intersectionf_rw=>// =>ip.
move: nep=> [j jp]; move: (ip _ jp)=> xp.
have aux: (forall i j, inc i (domain f) -> inc j (domain g) ->
(inc (V TPa x) (V i f) & inc (V TPb x) (V j g))).
move=> i k id kd.
have Jd: (inc (J i k) (product (domain f) (domain g))) by fprops.
move: (ip _ Jd); aw; move=> [fx [dx px]].
have ta: (inc TPa (domain x)) by ue.
have tb: (inc TPb (domain x)) by ue.
by move: (px _ ta)(px _ tb); bw; intuition.
move: xp; aw; move=> [fx [dx px]]; split=>//; split=>//.
move=> i; rewrite dx=> id.
try_lvariant id; apply intersectionb_inc; try rewrite -nonempty_domain //.
by move => k kd; move: ng=> [l lg]; move: (aux _ _ kd lg); intuition.
by move => l lg; move:nf=> [k kd]; move: (aux _ _ kd lg); intuition.
Qed.
Lemma distrib_product2_union: forall f g,
fgraph f -> fgraph g ->
product (unionb f)(unionb g) =
unionf(product (domain f)(domain g)) (fun z =>
(product (V (P z) f) (V (Q z) g))).
Proof.
move=> f g ff fg.
case (emptyset_dichot (domain f)) =>nf.
rewrite nf empty_unionf; srw.
by rewrite/unionb nf unionf_emptyset; srw.
case (emptyset_dichot (domain g)) => ng.
rewrite ng empty_unionf; srw.
by rewrite/unionb ng unionf_emptyset; srw.
move: (distrib_prod2_union ff fg nf ng); rewrite/product2 => distr.
set_extens x => xp.
move: xp;aw; move=> [xp [px qx]].
set (v := variantLc (P x) (Q x)).
have: inc v (productb (variantLc (unionb f) (unionb g))).
aw; fprops; rewrite/v; split; fprops; bw; split=>//.
by move=> i it; try_lvariant it.
rewrite /productf in distr; rewrite variantLc_prop distr.
srw; move=> [y [p1 p2]]; exists y; split=>//.
move: p2; aw; gprops; move=> [fv []]; bw => dv p.
have ta: (inc TPa (domain v)) by rewrite /v; bw; fprops.
have tb: (inc TPb (domain v)) by rewrite /v; bw; fprops.
by move: (p _ ta) (p _ tb); rewrite /v; bw; fprops.
move: xp; srw.
move=> [y [yp]]; rewrite 2! product_inc_rw; move=> [px [xp xq]].
set (qx := variantLc (P x) (Q x)).
have:(inc qx (productb (variantLc (unionb f) (unionb g)))).
move: distr; rewrite/productf -variantLc_prop; move=>->.
union_tac; aw; gprops;rewrite/ qx.
split; fprops; bw;split=>//.
move=> i it; try_lvariant it; fprops.
aw; fprops; rewrite variantLc_domain.
move=> [fq [dx vqx]].
have ta: (inc TPa (domain qx)) by rewrite dx; fprops.
have tb: (inc TPb (domain qx)) by rewrite dx; fprops.
move: (vqx _ ta) (vqx _ tb); rewrite /qx; bw.
intuition.
Qed.
Lemma distrib_product2_inter: forall f g,
fgraph f -> fgraph g -> nonempty (domain f) -> nonempty (domain g) ->
product (intersectionb f)(intersectionb g) =
intersectionf(product (domain f)(domain g)) (fun z =>
(product (V (P z) f) (V (Q z) g))).
Proof.
move=>f g ff fg nef neg.
move: (distrib_prod2_inter ff fg nef neg); rewrite / product2=> distr.
have nep: nonempty (product (domain f) (domain g)).
by move: nef neg; move=> [y yf][z zg]; exists (J y z); fprops.
set_extens x => xp.
set (qx := variantLc (P x) (Q x)).
have: inc qx (productb (variantLc (intersectionb f) (intersectionb g))).
move: xp; aw; gprops; fprops; move=>[xp [pxf qxg]].
rewrite /qx; split;fprops;bw; split=>//.
by move=> i ind; try_lvariant ind.
move: distr; rewrite /productf -variantLc_prop; move=> -> qxp.
apply intersectionf_inc=>//.
move=> j jp; move: (intersectionf_forall qxp jp).
move: xp; aw; gprops; move=> [px [xp xq]] [fgq [dq vq]].
move: dq; bw => dq; rewrite dq in vq.
have ta: (inc TPa two_points) by fprops.
have tb: (inc TPb two_points) by fprops.
move: (vq _ ta)(vq _ tb); rewrite /qx; bw.
by intuition.
set (qx := variantLc (P x) (Q x)).
have: inc qx (productb (variantLc (intersectionb f) (intersectionb g))).
move: distr; rewrite /productf variantLc_prop; move =>->.
apply intersectionf_inc=>//.
move=> j jp; move: (intersectionf_forall xp jp).
aw;gprops; move=> [_ [Px Qx]].
rewrite /qx; split;fprops; bw; split=>//.
move=> i it; try_lvariant it; fprops.
aw; fprops;move=> [_ [_]];rewrite /qx; bw => qv.
move: nep=> [w wp]; move: (intersectionf_forall xp wp); aw; move=> [px _].
split=>//.
move: (qv _ inc_TPa_two_points) (qv _ inc_TPb_two_points); bw.
intuition.
Qed.
A variant of distributivity, valid if the domain of the double famil
y is a rectangle; works only for intersection.
Theorem distrib_inter_prod: forall f sa sb,
fgraph f -> nonempty sb ->
intersectionf sb (fun k => productf sa (fun i=> V (J i k) f)) =
productf sa (fun i => intersectionf sb (fun k=> V (J i k) f)).
Proof.
move=> f sa sb ff nb; set_extens x => xi.
move: nb=> [y yb]; move: (intersectionf_forall xi yb).
aw; move => [fx [dx px]]; split=>//; split=>//.
move=> i id; apply intersectionf_inc; first by ex_tac.
by move => j jsb; move: (intersectionf_forall xi jsb); aw; intuition.
move: xi; aw; move=> [fx [dx px]].
apply intersectionf_inc =>//.
move=> j jsb; aw; split=>//; split=>//.
by move=> i id; move: (px _ id)=> xi; apply (intersectionf_forall xi jsb).
Qed.
Lemma productf_extension : forall sf1 f1 sf2 f2,
L sf1 f1 = L sf2 f2 -> productf sf1 f1 = productf sf2 f2.
Proof. by move=> sf1 gf1 sf2 f2 eq; rewrite /productf eq. Qed.
Lemma distrib_prod_inter2_prod: forall f g,
fgraph f -> fgraph g -> domain f = domain g ->
nonempty (domain f) ->
intersection2 (productb f) (productb g) =
productf (domain f) (fun i => intersection2 (V i f) (V i g)).
Proof.
move=> f g ff fg dfdg ne.
set_extens x; aw.
move=> [[fx [dx vx]][ _ [dx' vx']]].
split =>//; split=>//.
by move=> i id; apply intersection2_inc; [apply vx | apply vx'].
move=> [fx [dx vx]].
intuition; try ue; move: (vx _ H); aw; intuition.
Qed.
Lemma distrib_inter_prod_inter: forall f g,
fgraph f -> fgraph g -> domain f = domain g ->
nonempty (domain f) ->
product2 (intersectionb f) (intersectionb g) =
intersectionf (domain f) (fun i => product2 (V i f) (V i g)).
Proof.
move=> f g ff fg dfdg nedf.
have ta:inc TPa two_points by fprops.
have tb:inc TPb two_points by fprops.
set (sb:= domain f).
set (h:= L(product two_points sb)
(fun i => Yo (P i = TPa) (V (Q i) f) (V (Q i) g))).
have fh: fgraph h by rewrite /h; gprops.
have nes: nonempty sb by rewrite /sb.
move: (distrib_inter_prod two_points fh nes)=>distr.
have isb: intersectionf sb(fun k => productf two_points(fun i => V (J i k) h))
= intersectionf sb (fun i => product2 (V i f) (V i g)).
apply intersectionf_exten.
move=> i isb; rewrite /product2.
apply productf_extension; apply L_exten1=>//.
move => x xp; rewrite /h; try_lvariant xp; fprops ; aw.
by rewrite Y_if_rw //.
by rewrite Y_if_not_rw //; apply two_points_distinctb.
rewrite -isb distr /product2.
apply productf_extension;apply L_exten1=>//.
move=> x xb; rewrite /sb /intersectionb.
try_lvariant xb; try rewrite -dfdg; apply intersectionf_exten=>//.
move=> i id; rewrite /h;bw; fprops; aw; rewrite Y_if_rw //.
move=> i id; rewrite /h;bw; fprops; aw; rewrite Y_if_not_rw //.
apply two_points_distinctb.
Qed.
fgraph f -> nonempty sb ->
intersectionf sb (fun k => productf sa (fun i=> V (J i k) f)) =
productf sa (fun i => intersectionf sb (fun k=> V (J i k) f)).
Proof.
move=> f sa sb ff nb; set_extens x => xi.
move: nb=> [y yb]; move: (intersectionf_forall xi yb).
aw; move => [fx [dx px]]; split=>//; split=>//.
move=> i id; apply intersectionf_inc; first by ex_tac.
by move => j jsb; move: (intersectionf_forall xi jsb); aw; intuition.
move: xi; aw; move=> [fx [dx px]].
apply intersectionf_inc =>//.
move=> j jsb; aw; split=>//; split=>//.
by move=> i id; move: (px _ id)=> xi; apply (intersectionf_forall xi jsb).
Qed.
Lemma productf_extension : forall sf1 f1 sf2 f2,
L sf1 f1 = L sf2 f2 -> productf sf1 f1 = productf sf2 f2.
Proof. by move=> sf1 gf1 sf2 f2 eq; rewrite /productf eq. Qed.
Lemma distrib_prod_inter2_prod: forall f g,
fgraph f -> fgraph g -> domain f = domain g ->
nonempty (domain f) ->
intersection2 (productb f) (productb g) =
productf (domain f) (fun i => intersection2 (V i f) (V i g)).
Proof.
move=> f g ff fg dfdg ne.
set_extens x; aw.
move=> [[fx [dx vx]][ _ [dx' vx']]].
split =>//; split=>//.
by move=> i id; apply intersection2_inc; [apply vx | apply vx'].
move=> [fx [dx vx]].
intuition; try ue; move: (vx _ H); aw; intuition.
Qed.
Lemma distrib_inter_prod_inter: forall f g,
fgraph f -> fgraph g -> domain f = domain g ->
nonempty (domain f) ->
product2 (intersectionb f) (intersectionb g) =
intersectionf (domain f) (fun i => product2 (V i f) (V i g)).
Proof.
move=> f g ff fg dfdg nedf.
have ta:inc TPa two_points by fprops.
have tb:inc TPb two_points by fprops.
set (sb:= domain f).
set (h:= L(product two_points sb)
(fun i => Yo (P i = TPa) (V (Q i) f) (V (Q i) g))).
have fh: fgraph h by rewrite /h; gprops.
have nes: nonempty sb by rewrite /sb.
move: (distrib_inter_prod two_points fh nes)=>distr.
have isb: intersectionf sb(fun k => productf two_points(fun i => V (J i k) h))
= intersectionf sb (fun i => product2 (V i f) (V i g)).
apply intersectionf_exten.
move=> i isb; rewrite /product2.
apply productf_extension; apply L_exten1=>//.
move => x xp; rewrite /h; try_lvariant xp; fprops ; aw.
by rewrite Y_if_rw //.
by rewrite Y_if_not_rw //; apply two_points_distinctb.
rewrite -isb distr /product2.
apply productf_extension;apply L_exten1=>//.
move=> x xb; rewrite /sb /intersectionb.
try_lvariant xb; try rewrite -dfdg; apply intersectionf_exten=>//.
move=> i id; rewrite /h;bw; fprops; aw; rewrite Y_if_rw //.
move=> i id; rewrite /h;bw; fprops; aw; rewrite Y_if_not_rw //.
apply two_points_distinctb.
Qed.
Product of two products of family of sets
Definition prod_of_function x x':=
L(domain x)(fun i => J (V i x) (V i x')).
Definition prod_of_products_canon f f':=
BL(fun w => prod_of_function (P w) (Q w))
(product (productb f) (productb f'))
(productf (domain f)(fun i => product (V i f) (V i f'))).
Definition prod_of_product_aux f f' :=
fun i => (product (W i f) (W i f')).
Definition prod_of_prod_target f f' :=
fun_image(source f)(prod_of_product_aux f f').
Definition prod_of_products f f' :=
BL (prod_of_product_aux f f')(source f)(prod_of_prod_target f f').
Lemma prod_of_prod_inc_target: forall f f' x,
inc x (prod_of_prod_target f f') =
(exists i, inc i (source f) & product (W i f) (W i f') = x).
Proof. by rewrite /prod_of_prod_target=> f f' x; aw.
Qed.
Lemma prod_of_products_function: forall f f',
is_function (prod_of_products f f').
Proof.
rewrite /prod_of_products=> f f'; apply bl_function.
by move=> x xs; rewrite prod_of_prod_inc_target; ex_tac.
Qed.
Lemma prod_of_products_source: forall f f',
source (prod_of_products f f') = source f.
Proof. by rewrite /prod_of_products=> f f'; aw. Qed.
Lemma prod_of_products_target: forall f f',
target (prod_of_products f f') = prod_of_prod_target f f'.
Proof. by rewrite /prod_of_products=> f f'; aw. Qed.
Lemma prod_of_products_W: forall f f' i,
inc i (source f) ->
W i (prod_of_products f f') = product (W i f) (W i f').
Proof.
rewrite/prod_of_products=> f f' i isf; aw.
move=> x xf; rewrite prod_of_prod_inc_target; ex_tac.
Qed.
Lemma prod_of_products_fam_pr: forall f f' x,
is_function x -> source x = source f ->
target x = union (prod_of_prod_target f f') ->
inc (graph x) (productb (graph (prod_of_products f f'))) =
forall i, inc i (source f) ->
(is_pair (W i x) & inc (P (W i x)) (W i f) & inc (Q (W i x)) (W i f')).
Proof.
move=> f f' x fx sxsf txu.
transitivity( forall i, inc i (source f) ->
inc (V i (graph x)) (product (W i f) (W i f'))); last first.
apply iff_eq; rewrite /W; move=> eq i isf; move: (eq _ isf); aw.
have fp:is_function (prod_of_products f f') by apply prod_of_products_function.
have fg:fgraph (graph (prod_of_products f f')) by fprops.
apply iff_eq.
aw; move=> [fgf [sx iVV]] i isf.
move: (prod_of_products_W f' isf); rewrite {1} /W => aux.
by rewrite -sxsf in isf; move: (iVV _ isf); rewrite aux.
move=> hyp; aw; split; fprops.
split; first by rewrite /prod_of_products; aw.
rewrite sxsf; move=> i isx; move: (prod_of_products_W f' isx).
by rewrite {1} /W=> ->; apply hyp.
Qed.
Lemma prod_of_function_axioms:forall x x' f f',
is_function f -> is_function f' -> source f = source f' ->
inc (graph x) (productb (graph f)) -> inc (graph x') (productb (graph f')) ->
transf_axioms (fun i => J (W i x) (W i x'))
(source f) (union (prod_of_prod_target f f')).
Proof.
move=> x x' f f' ff ff' sfsf gp gp' y ys.
apply union_inc with (product (W y f) (W y f'));
last by rewrite prod_of_prod_inc_target; ex_tac.
move: gp gp'.
rewrite productb_rw;fprops; rewrite productb_rw;fprops.
move=>[fgx [dgx p1]][fg'x [dgx' p2]].
aw; split; fprops; split.
by apply p1; rewrite dgx f_domain_graph.
by apply p2; rewrite dgx' f_domain_graph // -sfsf.
Qed.
Lemma prod_of_function_W:forall x x' f f' i,
is_function f -> is_function f' -> source f = source f' ->
inc x (productb (graph f)) -> inc x' (productb (graph f')) ->
inc i (source f) ->
V i (prod_of_function x x') = J (V i x) (V i x').
Proof.
rewrite /prod_of_function=> x x' f f' i ff ff' sfsf xp x'p isf.
bw; move: xp; rewrite productb_rw; fprops; move=> [_ [dx _]].
by rewrite dx f_domain_graph.
Qed.
Lemma prod_of_function_function: forall x x' f f',
is_function f -> is_function f' -> source f = source f' ->
inc x (productb (graph f)) -> inc x' (productb (graph f')) ->
inc (prod_of_function x x')
(productb (graph (prod_of_products f f'))).
Proof.
move=> x x' f f' ff ff' sfsf xp x'p.
have fgf: fgraph (graph f) by fprops.
have fgf': fgraph (graph f') by fprops.
move: (prod_of_products_function f f') => fp.
rewrite /prod_of_function.
move: (xp); aw; fprops; move => [fgd [ dx ivv]].
aw;fprops; split; gprops;bw; split.
by rewrite prod_of_products_source.
move=> i idx; move: xp x'p; aw.
move=> [fx [sx ivvx]][fx' [sx' ivvx']].
bw; rewrite -[ (V i (graph _ ))]/ (W i (prod_of_products f f')).
rewrite prod_of_products_W; try ue.
by aw;split; fprops; split; [apply ivvx | apply ivvx'; rewrite sx' -sfsf -sx].
Qed.
Lemma popc_target_aux:forall f f',
is_function f -> is_function f' -> source f = source f' ->
productb(L (domain (graph f))
(fun i => product (V i (graph f)) (V i (graph f')))) =
productb (graph (prod_of_products f f')).
Proof.
move=> f f' ff ff' sfsf; apply f_equal.
have fp: (is_function (prod_of_products f f'))
by apply prod_of_products_function.
apply fgraph_exten;gprops;fprops; bw.
by rewrite /prod_of_products; aw.
move=> x xd; bw.
move: xd; rewrite (f_domain_graph ff) => xd.
by move: (prod_of_products_W f' xd);rewrite {1} /W; move=> ->.
Qed.
Lemma popc_target:forall f f',
is_function f -> is_function f' -> source f = source f' ->
target (prod_of_products_canon (graph f) (graph f')) =
(productb (graph (prod_of_products f f'))).
Proof.
rewrite /prod_of_products_canon/productf=> f f' ff ff' sfsf .
rewrite popc_target_aux //; aw.
Qed.
Lemma popc_axioms:forall f f',
is_function f -> is_function f' -> source f = source f' ->
transf_axioms(fun w => prod_of_function (P w) (Q w))
(product (productb (graph f)) (productb (graph f')))
(productb (graph (prod_of_products f f'))).
Proof.
move=> f f' ff sfsf x xd.
rewrite product_inc_rw.
move=> [pxd [px qx]]; apply prod_of_function_function =>//.
Qed.
Lemma popc_W:forall f f' w,
is_function f -> is_function f' -> source f = source f' ->
inc w (product (productb (graph f)) (productb (graph f'))) ->
W w (prod_of_products_canon (graph f) (graph f')) =
prod_of_function (P w) (Q w).
Proof.
move=> f f' w ff ff' sfsf h.
rewrite /prod_of_products_canon bl_W //.
rewrite /productf popc_target_aux //.
by apply popc_axioms.
Qed.
Lemma popc_bijection:forall f f',
is_function f -> is_function f' -> source f = source f' ->
bijection (prod_of_products_canon (graph f) (graph f')).
Proof.
move=> f f' ff ff' sfsf.
set (g:= prod_of_products_canon (graph f) (graph f')).
have sg: source g = product (productb (graph f)) (productb (graph f')).
by rewrite /g/prod_of_products_canon; aw.
have fg: (is_function g).
rewrite /g/prod_of_products_canon; apply bl_function.
move=> x xs; rewrite/productf popc_target_aux //.
by apply popc_axioms=>//.
split.
split=>//.
move=> x y; rewrite sg=> xs ys.
rewrite popc_W // popc_W // => sv.
move: xs ys ; rewrite 2! product_inc_rw; move=> [xp [px qx]][yp [py qy]].
have sj: (forall i, inc i (source f) ->
J (V i (P x)) (V i (Q x)) = J (V i (P y)) (V i (Q y))).
move=> i isf.
rewrite - (prod_of_function_W (f:=f)(f':=f')) //.
rewrite - (prod_of_function_W (f:=f)(f':=f')) //.
by rewrite sv.
apply pair_exten=>//.
apply productb_exten with (f:=(graph f)) =>//; fprops.
move=> j; rewrite (f_domain_graph ff) => js.
move: (sj _ js); apply pr1_def.
apply productb_exten with (f:=(graph f')) =>//; fprops.
move=> j; rewrite (f_domain_graph ff') -sfsf => js.
by move: (sj _ js); apply pr2_def.
have fp: (is_function (prod_of_products f f')).
by apply prod_of_products_function.
apply surjective_pr6 =>//.
rewrite/g; rewrite popc_target // sg.
move=>y; aw; fprops; move=> [fy [dy iVV]].
rewrite /prod_of_products in dy; awi dy.
have iww : forall i : Set,
inc i (domain y) -> inc (V i y) (W i (prod_of_products f f')).
by move => i; apply iVV.
set (xa:= BL(fun i => P(V i y)) (source f)(union (target f))).
set (xb:= BL(fun i => Q(V i y)) (source f)(union (target f'))).
have ta: transf_axioms (fun i => P (V i y)) (source f) (union (target f)).
move=> t tf;rewrite dy in iww.
move: (iww _ tf); move: (prod_of_products_W f' tf); move=> ->.
aw; move=> [_ [ iP _]].
by apply union_inc with (V t (graph f)) => //; apply inc_W_target => //.
have tb: transf_axioms (fun i => Q (V i y)) (source f) (union (target f')).
move=> t tf;rewrite dy in iww.
move: (iww _ tf); move: (prod_of_products_W f' tf); move=> ->.
aw; move=> [_ [_ iQ]].
apply union_inc with (V t (graph f')) => //; apply inc_W_target => //.
by rewrite -sfsf.
have fxa: is_function xa by rewrite /xa; apply bl_function.
have fxb: is_function xb by rewrite /xb; apply bl_function.
have sxa: source xa = source f by rewrite /xa; aw.
have sxb: source xb = source f by rewrite /xb; aw.
have ixa :inc (graph xa) (productb (graph f)).
aw; fprops; split; fprops; split=>//.
rewrite sxa; move=> i isa; change (inc (W i xa) (W i f)).
rewrite /xa; aw.
move: (prod_of_products_W f' isa)=> eq.
by rewrite -dy in isa; move: (iww _ isa); rewrite eq; aw; intuition.
have ixb :inc (graph xb) (productb (graph f')).
aw; fprops; split; fprops; split; try ue.
rewrite sxb; move=> i isb; change (inc (W i xb) (W i f')).
rewrite /xb; aw.
move: (prod_of_products_W f' isb) => eq.
by rewrite -dy in isb; move: (iww _ isb); rewrite eq; aw; intuition.
exists (J (graph xa) (graph xb)).
split; fprops.
rewrite /g popc_W =>//; fprops; aw.
rewrite /prod_of_function.
have dga: domain (graph xa) = domain y by rewrite /xa; aw; ue.
apply fgraph_exten;gprops; bw; rewrite dga.
move => x xd; bw; rewrite dy in xd.
change (J (W x xa) (W x xb) = V x y).
rewrite /xa/xb; aw.
move: (prod_of_products_W f' xd) =>eq; rewrite -dy in xd; move: (iww _ xd).
by rewrite eq; aw; intuition.
Qed.
Definition ext_map_prod_aux x f := fun i=> V (V i x) (f i).
Definition ext_map_prod I src trg f :=
BL (fun x => L I (ext_map_prod_aux x f))
(productf I src ) (productf I trg ).
Definition ext_map_prod_axioms I src trg f :=
forall i, inc i I -> (fgraph (f i) & domain (f i) = src i &
sub (range (f i)) (trg i)).
Lemma ext_map_prod_taxioms: forall I src trg f,
ext_map_prod_axioms I src trg f ->
transf_axioms (fun x => L I (ext_map_prod_aux x f))
(productf I src) (productf I trg).
Proof.
move=> I src trg f ax x xs; aw; split; gprops; bw; split =>//.
move=> i iI; bw.
rewrite /ext_map_prod_aux.
move : (ax _ iI) => [fi [dfi sr]]; apply sr.
move: xs; rewrite productf_rw;move=> [_ [dx ivs]].
apply inc_V_range=>//; rewrite dfi; apply ivs; ue.
Qed.
Lemma ext_map_prod_function: forall I src trg f,
ext_map_prod_axioms I src trg f -> is_function ( ext_map_prod I src trg f).
Proof.
rewrite /ext_map_prod=> I src trg f hyp.
by apply bl_function; apply ext_map_prod_taxioms.
Qed.
Lemma ext_map_prod_W: forall I src trg f x,
ext_map_prod_axioms I src trg f ->
inc x (productf I src) ->
W x (ext_map_prod I src trg f) = L I (ext_map_prod_aux x f).
Proof.
rewrite /ext_map_prod=> I src trg f x hyp xp.
by aw; apply ext_map_prod_taxioms.
Qed.
Lemma ext_map_prod_WV: forall I src trg f x i,
ext_map_prod_axioms I src trg f ->
inc x (productf I src) -> inc i I ->
V i (W x (ext_map_prod I src trg f)) = V (V i x) (f i).
Proof.
move=> I src trg f x i ax xp iI.
rewrite ext_map_prod_W //; bw.
Qed.
Lemma ext_map_prod_composable: forall I p1 p2 p3 g f h,
ext_map_prod_axioms I p1 p2 f ->
ext_map_prod_axioms I p2 p3 g ->
(forall i, inc i I -> h i = fcompose (g i) (f i)) ->
(forall i, inc i I -> fcomposable (g i) (f i)) ->
ext_map_prod_axioms I p1 p3 h.
Proof.
move=> I p1 p2 p3 f g h emp1 emp2 alh alfc x xs.
move: (emp1 _ xs) (emp2 _ xs) => [fgg [dg srg]] [fgf [df srf]].
move: (alfc _ xs)(alh _ xs) => fca ->.
have fgc: fgraph (fcompose (f x) (g x)) by apply fcompose_fgraph.
have dfc :domain (fcompose (f x) (g x)) = domain (g x)
by rewrite fcomposable_domain.
rewrite dfc dg; split=>//; split=>//.
move=> y; srw.
move=> [z [zdg]] ->;rewrite fcompose_ev //; last by ue.
apply srf; srw; exists (V z (g x)); split=> //.
rewrite df; apply srg; apply inc_V_range =>//; ue.
Qed.
Lemma ext_map_prod_compose: forall I p1 p2 p3 g f h,
ext_map_prod_axioms I p1 p2 f ->
ext_map_prod_axioms I p2 p3 g ->
(forall i, inc i I -> h i = fcompose (g i) (f i)) ->
(forall i, inc i I -> fcomposable (g i) (f i)) ->
compose (ext_map_prod I p2 p3 g) (ext_map_prod I p1 p2 f) =
(ext_map_prod I p1 p3 h).
Proof.
move=> I p1 p2 p3 g f h emp1 emp2 alh alfc.
set f1:= (ext_map_prod I p2 p3 g).
set f2:= (ext_map_prod I p1 p2 f).
set f3:= (ext_map_prod I p1 p3 h).
move: (ext_map_prod_composable emp1 emp2 alh alfc) => emp3.
have cee: (f1 \coP f2).
split; first by apply ext_map_prod_function.
split; first by apply ext_map_prod_function.
rewrite /f1/f2/ext_map_prod; aw.
have fcee: (is_function (f1 \co f2)) by fct_tac.
apply function_exten=>//; try solve [ rewrite/f1/f2/f3/ext_map_prod; aw].
by apply ext_map_prod_function =>//.
aw; move=> x xs; aw.
have xp: inc x (productf I p1) by move: xs; rewrite /f2/ext_map_prod; aw.
set (t:= W x f2).
have tv: t = L I (ext_map_prod_aux x f) by rewrite /t/f2 ext_map_prod_W //.
have: (W t f1 = L I (ext_map_prod_aux t g)).
rewrite /f1 ext_map_prod_W //.
have tf2: (target f2 = productf I p2) by rewrite/f2/ext_map_prod; aw.
rewrite/t -tf2; apply inc_W_target =>//.
by rewrite/f2; apply ext_map_prod_function =>//.
have: (W x f3 = L I (ext_map_prod_aux x h)) by rewrite/f3 ext_map_prod_W //.
move=> ->->.
apply L_exten1=>//.
move => e ei; rewrite tv /ext_map_prod_aux; bw.
move: (alfc _ ei)=> aux.
rewrite (alh _ ei) fcompose_ev //.
move: (emp1 _ ei)=> [fgg [dg srg]]; rewrite dg.
move: xp; aw; move=> [_ [dx]]; apply; ue.
Qed.
Lemma ext_map_prod_injective: forall I p1 p2 f,
ext_map_prod_axioms I p1 p2 f ->
(forall i, inc i I -> injective_graph (f i)) ->
injection (ext_map_prod I p1 p2 f).
Proof.
move=> I p1 p2 f emp alli; split.
by apply ext_map_prod_function.
rewrite /ext_map_prod; rewrite bl_source; move => x y xp yp eql.
apply (productf_exten xp yp).
move=>i iI.
move: (ext_map_prod_WV emp xp iI); move: (ext_map_prod_WV emp yp iI).
rewrite eql; move=>-> aux.
move: (emp _ iI)=> [_ [df sr]].
move: (alli _ iI); move=> [ff tmp]; apply tmp=>//.
by rewrite df; move: xp; aw; move=> [_ [dx]]; apply; ue.
rewrite df; move: yp; aw; move=> [_ [dx]]; apply; ue.
Qed.
Lemma ext_map_prod_surjective: forall I p1 p2 f,
ext_map_prod_axioms I p1 p2 f ->
(forall i, inc i I -> range (f i) = p2 i) ->
surjection (ext_map_prod I p1 p2 f).
Proof.
move=> I p1 p2 f emp als;
set g :=(ext_map_prod I p1 p2 f).
have fg: is_function g by apply ext_map_prod_function.
apply surjective_pr6=>//.
have semp: source g = productf I p1 by rewrite /g /ext_map_prod bl_source.
have temp: target g = productf I p2 by rewrite /g /ext_map_prod bl_target.
rewrite semp temp; move=> y yt.
have ext: (forall i, inc i I -> exists z, inc z (p1 i) & V i y =V z (f i)).
move=> i iI; move: (emp _ iI) => [ff [df sr]].
move: yt; aw; move=> [fy [dy ip2]].
rewrite dy in ip2; move: (ip2 _ iI).
by rewrite - (als _ iI); srw; rewrite df.
set (x:= L I (fun i => choose (fun z => inc z (p1 i) & V i y = V z (f i)))).
have xp: forall i, inc i I -> (inc (V i x) (p1 i) & V i y = V (V i x) (f i)).
by move=> i iI; rewrite/x; bw; apply choose_pr; apply ext.
have ixp: (inc x (productf I p1)).
aw; rewrite /x; split; gprops; bw;split =>//; rewrite -/x.
by move=> i iI; bw; move: (xp _ iI); intuition.
have xsq: (inc x (source g)) by ue.
have iWp: (inc (W x g)(productf I p2)) by rewrite -temp; apply inc_W_target.
ex_tac.
apply (productf_exten iWp yt).
move=> i iI; rewrite ext_map_prod_WV //.
move: (xp _ iI); intuition.
Qed.
Lemma fun_set_to_prod1: forall F f i,
fgraph F -> inc i (domain F) ->
is_function f -> target f = productb F ->
(is_function (pr_i F i \co f) &
source (pr_i F i \co f) = source f &
target (pr_i F i \co f) = V i F&
(forall x, inc x (source f) -> W x (pr_i F i \co f) = V i (W x f))).
Proof.
move=> F f i fF id ff tfpF.
have aux: (pr_i F i) \coP f.
split; first by apply pri_function=>//.
by split =>//; symmetry; rewrite /pr_i; aw.
split; first by fct_tac.
split; first rewrite/pr_i; aw.
split; first rewrite/pr_i; aw.
by move=> x xs; aw; rewrite pri_W //; rewrite -tfpF; fprops.
Qed.
Definition ext_map_prod I src trg f :=
BL (fun x => L I (ext_map_prod_aux x f))
(productf I src ) (productf I trg ).
Definition ext_map_prod_axioms I src trg f :=
forall i, inc i I -> (fgraph (f i) & domain (f i) = src i &
sub (range (f i)) (trg i)).
Lemma ext_map_prod_taxioms: forall I src trg f,
ext_map_prod_axioms I src trg f ->
transf_axioms (fun x => L I (ext_map_prod_aux x f))
(productf I src) (productf I trg).
Proof.
move=> I src trg f ax x xs; aw; split; gprops; bw; split =>//.
move=> i iI; bw.
rewrite /ext_map_prod_aux.
move : (ax _ iI) => [fi [dfi sr]]; apply sr.
move: xs; rewrite productf_rw;move=> [_ [dx ivs]].
apply inc_V_range=>//; rewrite dfi; apply ivs; ue.
Qed.
Lemma ext_map_prod_function: forall I src trg f,
ext_map_prod_axioms I src trg f -> is_function ( ext_map_prod I src trg f).
Proof.
rewrite /ext_map_prod=> I src trg f hyp.
by apply bl_function; apply ext_map_prod_taxioms.
Qed.
Lemma ext_map_prod_W: forall I src trg f x,
ext_map_prod_axioms I src trg f ->
inc x (productf I src) ->
W x (ext_map_prod I src trg f) = L I (ext_map_prod_aux x f).
Proof.
rewrite /ext_map_prod=> I src trg f x hyp xp.
by aw; apply ext_map_prod_taxioms.
Qed.
Lemma ext_map_prod_WV: forall I src trg f x i,
ext_map_prod_axioms I src trg f ->
inc x (productf I src) -> inc i I ->
V i (W x (ext_map_prod I src trg f)) = V (V i x) (f i).
Proof.
move=> I src trg f x i ax xp iI.
rewrite ext_map_prod_W //; bw.
Qed.
Lemma ext_map_prod_composable: forall I p1 p2 p3 g f h,
ext_map_prod_axioms I p1 p2 f ->
ext_map_prod_axioms I p2 p3 g ->
(forall i, inc i I -> h i = fcompose (g i) (f i)) ->
(forall i, inc i I -> fcomposable (g i) (f i)) ->
ext_map_prod_axioms I p1 p3 h.
Proof.
move=> I p1 p2 p3 f g h emp1 emp2 alh alfc x xs.
move: (emp1 _ xs) (emp2 _ xs) => [fgg [dg srg]] [fgf [df srf]].
move: (alfc _ xs)(alh _ xs) => fca ->.
have fgc: fgraph (fcompose (f x) (g x)) by apply fcompose_fgraph.
have dfc :domain (fcompose (f x) (g x)) = domain (g x)
by rewrite fcomposable_domain.
rewrite dfc dg; split=>//; split=>//.
move=> y; srw.
move=> [z [zdg]] ->;rewrite fcompose_ev //; last by ue.
apply srf; srw; exists (V z (g x)); split=> //.
rewrite df; apply srg; apply inc_V_range =>//; ue.
Qed.
Lemma ext_map_prod_compose: forall I p1 p2 p3 g f h,
ext_map_prod_axioms I p1 p2 f ->
ext_map_prod_axioms I p2 p3 g ->
(forall i, inc i I -> h i = fcompose (g i) (f i)) ->
(forall i, inc i I -> fcomposable (g i) (f i)) ->
compose (ext_map_prod I p2 p3 g) (ext_map_prod I p1 p2 f) =
(ext_map_prod I p1 p3 h).
Proof.
move=> I p1 p2 p3 g f h emp1 emp2 alh alfc.
set f1:= (ext_map_prod I p2 p3 g).
set f2:= (ext_map_prod I p1 p2 f).
set f3:= (ext_map_prod I p1 p3 h).
move: (ext_map_prod_composable emp1 emp2 alh alfc) => emp3.
have cee: (f1 \coP f2).
split; first by apply ext_map_prod_function.
split; first by apply ext_map_prod_function.
rewrite /f1/f2/ext_map_prod; aw.
have fcee: (is_function (f1 \co f2)) by fct_tac.
apply function_exten=>//; try solve [ rewrite/f1/f2/f3/ext_map_prod; aw].
by apply ext_map_prod_function =>//.
aw; move=> x xs; aw.
have xp: inc x (productf I p1) by move: xs; rewrite /f2/ext_map_prod; aw.
set (t:= W x f2).
have tv: t = L I (ext_map_prod_aux x f) by rewrite /t/f2 ext_map_prod_W //.
have: (W t f1 = L I (ext_map_prod_aux t g)).
rewrite /f1 ext_map_prod_W //.
have tf2: (target f2 = productf I p2) by rewrite/f2/ext_map_prod; aw.
rewrite/t -tf2; apply inc_W_target =>//.
by rewrite/f2; apply ext_map_prod_function =>//.
have: (W x f3 = L I (ext_map_prod_aux x h)) by rewrite/f3 ext_map_prod_W //.
move=> ->->.
apply L_exten1=>//.
move => e ei; rewrite tv /ext_map_prod_aux; bw.
move: (alfc _ ei)=> aux.
rewrite (alh _ ei) fcompose_ev //.
move: (emp1 _ ei)=> [fgg [dg srg]]; rewrite dg.
move: xp; aw; move=> [_ [dx]]; apply; ue.
Qed.
Lemma ext_map_prod_injective: forall I p1 p2 f,
ext_map_prod_axioms I p1 p2 f ->
(forall i, inc i I -> injective_graph (f i)) ->
injection (ext_map_prod I p1 p2 f).
Proof.
move=> I p1 p2 f emp alli; split.
by apply ext_map_prod_function.
rewrite /ext_map_prod; rewrite bl_source; move => x y xp yp eql.
apply (productf_exten xp yp).
move=>i iI.
move: (ext_map_prod_WV emp xp iI); move: (ext_map_prod_WV emp yp iI).
rewrite eql; move=>-> aux.
move: (emp _ iI)=> [_ [df sr]].
move: (alli _ iI); move=> [ff tmp]; apply tmp=>//.
by rewrite df; move: xp; aw; move=> [_ [dx]]; apply; ue.
rewrite df; move: yp; aw; move=> [_ [dx]]; apply; ue.
Qed.
Lemma ext_map_prod_surjective: forall I p1 p2 f,
ext_map_prod_axioms I p1 p2 f ->
(forall i, inc i I -> range (f i) = p2 i) ->
surjection (ext_map_prod I p1 p2 f).
Proof.
move=> I p1 p2 f emp als;
set g :=(ext_map_prod I p1 p2 f).
have fg: is_function g by apply ext_map_prod_function.
apply surjective_pr6=>//.
have semp: source g = productf I p1 by rewrite /g /ext_map_prod bl_source.
have temp: target g = productf I p2 by rewrite /g /ext_map_prod bl_target.
rewrite semp temp; move=> y yt.
have ext: (forall i, inc i I -> exists z, inc z (p1 i) & V i y =V z (f i)).
move=> i iI; move: (emp _ iI) => [ff [df sr]].
move: yt; aw; move=> [fy [dy ip2]].
rewrite dy in ip2; move: (ip2 _ iI).
by rewrite - (als _ iI); srw; rewrite df.
set (x:= L I (fun i => choose (fun z => inc z (p1 i) & V i y = V z (f i)))).
have xp: forall i, inc i I -> (inc (V i x) (p1 i) & V i y = V (V i x) (f i)).
by move=> i iI; rewrite/x; bw; apply choose_pr; apply ext.
have ixp: (inc x (productf I p1)).
aw; rewrite /x; split; gprops; bw;split =>//; rewrite -/x.
by move=> i iI; bw; move: (xp _ iI); intuition.
have xsq: (inc x (source g)) by ue.
have iWp: (inc (W x g)(productf I p2)) by rewrite -temp; apply inc_W_target.
ex_tac.
apply (productf_exten iWp yt).
move=> i iI; rewrite ext_map_prod_WV //.
move: (xp _ iI); intuition.
Qed.
Lemma fun_set_to_prod1: forall F f i,
fgraph F -> inc i (domain F) ->
is_function f -> target f = productb F ->
(is_function (pr_i F i \co f) &
source (pr_i F i \co f) = source f &
target (pr_i F i \co f) = V i F&
(forall x, inc x (source f) -> W x (pr_i F i \co f) = V i (W x f))).
Proof.
move=> F f i fF id ff tfpF.
have aux: (pr_i F i) \coP f.
split; first by apply pri_function=>//.
by split =>//; symmetry; rewrite /pr_i; aw.
split; first by fct_tac.
split; first rewrite/pr_i; aw.
split; first rewrite/pr_i; aw.
by move=> x xs; aw; rewrite pri_W //; rewrite -tfpF; fprops.
Qed.
The two sets (prod Fi)^E and prod (Fi^E) are isomorphic
Definition fun_set_to_prod src F :=
BL(fun f =>
L(domain F)( fun i=> (graph ( (pr_i F i) \co
(corresp src (productb F) f)))))
(set_of_gfunctions src (productb F))
(productb (L (domain F) (fun i=> set_of_gfunctions src (V i F)))).
Lemma fun_set_to_prod2: forall src F f gf,
fgraph F -> inc gf (set_of_gfunctions src (productb F)) ->
f = (corresp src (productb F) gf) -> function_prop f src (productb F).
Proof.
move=> src F f gf fF igf eq.
move: (set_of_gfunctions_inc igf) => [g [fg [sg [tg gg]]]].
have : f = g.
by rewrite eq -sg -tg -gg; apply corresp_recov1; move: fg=> [ res _].
move=> ->; red; intuition.
Qed.
Lemma fun_set_to_prod3: forall src F,
fgraph F -> transf_axioms(fun f =>
L(domain F)( fun i=> (graph (compose (pr_i F i)
(corresp src (productb F) f)))))
(set_of_gfunctions src (productb F))
(productb (L (domain F) (fun i=> set_of_gfunctions src (V i F)))).
Proof.
move=> src F fF x xs.
set (f:= corresp src (productb F) x).
move: (fun_set_to_prod2 fF xs (refl_equal f)) => [ff [sf tf]].
aw; gprops; split ;gprops; bw; split=>//.
move=> i id; bw.
move: (fun_set_to_prod1 fF id ff tf) => [fc [sc [tg vc]]].
by rewrite -sf -sc -tg; apply inc_set_of_gfunctions; gprops.
Qed.
Lemma fun_set_to_prod4: forall src F, fgraph F ->
function_prop (fun_set_to_prod src F) (set_of_gfunctions src (productb F))
(productb (L (domain F) (fun i=> set_of_gfunctions src (V i F)))).
Proof.
rewrite/fun_set_to_prod=>src F FF; red;aw;intuition.
by apply bl_function; apply fun_set_to_prod3.
Qed.
Definition fun_set_to_prod5 F f :=
ext_map_prod (domain F) (fun i=> source f)(fun i=> V i F)
(fun i => (graph (compose (pr_i F i) f))).
Lemma fun_set_to_prod6: forall F f,
fgraph F -> is_function f -> target f = productb F ->
(is_function (fun_set_to_prod5 F f) &
(fun_set_to_prod5 F f) \coP (constant_functor (domain F)(source f) )&
(fun_set_to_prod5 F f) \co (constant_functor (domain F)(source f)) =f).
Proof.
move=> F f fF ff tfpF.
set (tf := fun_set_to_prod5 F f).
have ftf: is_function tf.
rewrite /tf/fun_set_to_prod5; apply ext_map_prod_function.
move=> x xs.
move: (fun_set_to_prod1 fF xs ff tfpF) => [fc [sc [tc vc]]].
by split; fprops; aw; split=>//; rewrite -tc; apply f_range_graph.
have ttf:(productb F = target tf).
rewrite /tf/fun_set_to_prod5/ext_map_prod.
aw; rewrite /productf; apply f_equal; apply fgraph_exten; gprops; bw.
by move=> x xd; bw.
have ctf: tf \coP (constant_functor (domain F)(source f)).
split; fprops; split.
rewrite /constant_functor; apply bl_function.
move => x xs; rewrite graphset_pr2.
rewrite/cst_graph; bw; intuition; gprops.
by move=> y; bw; move=> [b [bf]] <-.
set (k:= cst_graph (domain F) (source f)).
have: (domain F = domain k) by rewrite/k /cst_graph; bw.
move=> ->; rewrite /constant_functor; aw.
rewrite - product_eq_graphset.
by rewrite /tf/fun_set_to_prod5/ext_map_prod; aw.
by rewrite /k /cst_graph; gprops.
by move => i; rewrite /k /cst_graph; rewrite L_domain=> id; bw.
split=> //; split=>//.
have ta1: transf_axioms [eta cst_graph (domain F)]
(source f) (set_of_gfunctions (domain F) (source f)).
move=> y yd; rewrite /constant_graph/set_of_gfunctions.
apply Z_inc; last by rewrite /cst_graph; bw; split; gprops.
apply powerset_inc; move=> t; rewrite /L fun_image_rw.
by move=> [z [zd]] <-; fprops.
symmetry; apply function_exten =>//; try fct_tac.
by rewrite /constant_functor; aw.
rewrite /tf/fun_set_to_prod5/ext_map_prod; aw.
by rewrite /productf tfpF L_recovers //.
move=> x xsf; rewrite /constant_functor; aw.
have cgd: inc (cst_graph (domain F) x)
(productf (domain F) (fun _ : Set => source f)).
rewrite/cst_graph; aw.
by split; gprops; bw; split; bw; move => j jd; bw.
have:inc (W x f) (productb F) by rewrite - tfpF; fprops.
have: (inc (W (cst_graph (domain F) x) tf) (productb F)).
rewrite ttf; apply inc_W_target=>//.
rewrite /tf/fun_set_to_prod5/ext_map_prod; aw.
rewrite /cst_graph.
by split;gprops; bw; split=>//; move => i id; bw.
move=> p1 p2; apply (productb_exten fF p2 p1).
move=> i idf; rewrite /tf/fun_set_to_prod5 ext_map_prod_W //.
bw; rewrite /ext_map_prod_aux /cst_graph; bw.
set cp:= (pr_i F i) \co f.
rewrite -[V x (graph cp)]/(W x cp) /cp;aw.
rewrite pri_W //.
split.
by apply pri_function=>//; split=>//.
by rewrite /pr_i; aw.
move => j jd.
move: (fun_set_to_prod1 fF jd ff tfpF) => [fcf [sc [tc vc]]].
split; fprops; split.
by rewrite -sc; fprops; aw.
by rewrite -tc;apply f_range_graph=>//.
Qed.
Lemma fun_set_to_prod7: forall src F f g ,
fgraph F ->
(forall i, inc i (domain F) -> inc (f i) (set_of_gfunctions src (V i F))) ->
g = ext_map_prod (domain F) (fun i=> src)(fun i=> V i F) f ->
(forall i, inc i (domain F) ->
f i = graph ((pr_i F i) \co (g \co (constant_functor (domain F) src) ))).
Proof.
move=> src F f g fF allf geq i idf.
set f1:= (constant_functor (domain F) src).
have fct: is_function f1.
rewrite /f1 /constant_functor; apply bl_function.
move=> x xd;rewrite graphset_pr2 /cst_graph.
by split; gprops ; bw; split =>//;move => w; bw; move=> [b [_] ] <-.
have fg: (is_function g).
rewrite geq; apply ext_map_prod_function.
move=> x xd; move:(set_of_gfunctions_inc (allf _ xd))=> [h [? [? [th gh]]]].
rewrite -gh f_domain_graph //; split; fprops; split=>//.
by rewrite -th; apply f_range_graph=>//.
have cpa: g \coP f1.
split=>//; split=>//.
rewrite geq/f1/ext_map_prod/constant_functor; aw.
rewrite /productf.
set (k:= (L (domain F) (fun _ : Set => src))).
have dkdf: (domain k = domain F) by rewrite /k;bw.
have fgk: (fgraph k) by rewrite /k; gprops.
have aux: (forall i, inc i(domain k) -> V i k = src).
by move=> j jd; rewrite /k; bw;ue.
by rewrite -dkdf; apply (product_eq_graphset fgk aux).
have fcf: is_function (g \co f1) by fct_tac.
have cpb: (pr_i F i) \coP (g \co f1).
split; first by apply pri_function => //.
split=>//; rewrite geq.
by rewrite /pr_i/ext_map_prod; aw; rewrite/productf L_recovers //.
have fctr: is_function ((pr_i F i) \co (g \co f1)) by fct_tac.
move: (allf _ idf) => aux; move: (set_of_gfunctions_inc aux).
move=> [h [fh [sh [th gh]]]].
have sf1: source f1 = source h by rewrite /f1/pr_i/constant_functor; aw.
have ta1: transf_axioms (fun x0 : Set => cst_graph (domain F) x0)
src (set_of_gfunctions (domain F) src).
move=> y yd; rewrite /cst_graph/set_of_gfunctions.
apply Z_inc; last by bw; split; gprops.
apply powerset_inc; move=> t; rewrite /L fun_image_rw.
by move=> [z [zd]] <-; fprops.
suff: (h = (pr_i F i) \co (g \co f1)) by move=> <-; ue.
apply function_exten =>//.
by rewrite /pr_i/constant_functor; aw.
by rewrite /pr_i/constant_functor; aw.
rewrite /f1; move => x xh.
have xs: inc x src by rewrite -sh.
have xs2: inc x (source (constant_functor (domain F) src)).
by rewrite /constant_functor; aw.
have aax: ext_map_prod_axioms (domain F) (fun _ : Set => src)
(fun i0 : Set => V i0 F) f.
move=> z zs; move: (allf _ zs) => p;move: (set_of_gfunctions_inc p).
move=> [k [fk [sk [tk gk]]]].
rewrite -gk -sk -tk f_domain_graph //; intuition.
by apply f_range_graph=>//.
set q:= cst_graph (domain F) x.
have qp: inc q (productf (domain F) (fun _ : Set => src)).
aw; rewrite/q /cst_graph; bw; split; gprops; split =>//.
by move=> k kd; bw.
have si: W x (constant_functor (domain F) src) = q.
by rewrite /W/constant_functor /BL; aw; bw.
have seq: (x = V i q) by rewrite constant_graph_V //.
aw; rewrite si geq ext_map_prod_W //.
set s:= (L (domain F) (ext_map_prod_aux q f)).
have isp: inc s (productb F).
rewrite /s /ext_map_prod_aux.
aw; split; gprops; bw; split =>//.
move=> j jd; bw; rewrite constant_graph_V //.
move: (allf _ jd) => p;move: (set_of_gfunctions_inc p).
by move=> [k [fk [sk [tk gk]]]]; rewrite -gk -tk; apply inc_W_target=>//; ue.
aw; rewrite pri_W => //.
by rewrite /s; bw; rewrite /ext_map_prod_aux /W gh -seq.
Qed.
Lemma fun_set_to_prod8: forall src F,
fgraph F -> bijection (fun_set_to_prod src F).
Proof.
move=> src F fF.
move: (fun_set_to_prod4 src fF).
set (g :=fun_set_to_prod src F) in *.
move => [fg [sg tg]].
split.
split =>//; move=> x y; rewrite sg => xs ys.
move :(set_of_gfunctions_inc xs) (set_of_gfunctions_inc ys).
move=> [x0 [fx0 [sx0 [tx0 gx]]]][x1 [fx1 [sx1 [tx1 gx1]]]] eq.
suff: x0 = x1 by rewrite - gx -gx1; move=>->.
move:(fun_set_to_prod3 (src:=src) fF)=> ta.
move: eq; rewrite /g/fun_set_to_prod; aw; clear ta; move => eq.
apply function_exten=>//; try ue.
move=> x2 x2s.
have iW0: (inc (W x2 x0) (productb F)) by ue.
have iW1: (inc (W x2 x1) (productb F)) by rewrite sx0 -sx1 in x2s; ue.
apply (productb_exten fF iW0 iW1).
move=> i iF.
move: (fun_set_to_prod1 fF iF fx0 tx0) (fun_set_to_prod1 fF iF fx1 tx1).
move=> [fc0 [sc0 [tc0 vc0]]] [fc1 [sc1 [tc1 vc1]]].
have x0c: (x0 = corresp src (productb F) x).
by rewrite -sx0 -tx0 -gx corresp_recov1 //; move: fx0=>[cx0 _].
have x1c: (x1 = corresp src (productb F) y).
by rewrite -sx1 -tx1 -gx1 corresp_recov1 //; move: fx1=>[cx1 _].
have cc: ((pr_i F i) \co x0 = (pr_i F i) \co x1).
move: (f_equal (V i) eq); bw; rewrite - x0c -x1c => eq1.
by apply function_exten1 =>//; aw.
by rewrite - vc0 //; rewrite sx0 -sx1 in x2s; rewrite cc - vc1 //.
apply surjective_pr6 => // y.
rewrite sg {1} /g /fun_set_to_prod bl_target; aw; gprops.
move=> [fy]; bw; move=> [dy py].
set (x:= BL (fun u=> L (domain F) (fun i=> V u (V i y))) src (productb F)).
have sx: source x = src by rewrite /x; aw.
have tx: target x = productb F by rewrite /x; aw.
have fx: is_function x.
rewrite /x; apply bl_function.
move=> t ts; aw; split; gprops; bw; split =>//.
move => i idf; bw.
rewrite dy in py; move: (py _ idf); bw=> xpp.
move: (set_of_gfunctions_inc xpp) => [k [fk [sk [tk gk]]]].
have: (V t (V i y) = W t k) by rewrite /W gk.
by move=>->; rewrite -tk; apply inc_W_target =>//; ue.
have gs: inc (graph x) (set_of_gfunctions src (productb F)).
rewrite -sx -tx; apply inc_set_of_gfunctions=> //.
ex_tac.
move: (fun_set_to_prod3 fF (src:=src)) => ax.
rewrite /g/fun_set_to_prod; aw; clear ax.
apply fgraph_exten; bw;gprops =>//.
move=> z zd; bw.
have zdd: inc z (domain y) by ue.
move: (py _ zdd); bw; move => Vzg.
have xp: (x = corresp src (productb F) (graph x)).
by rewrite -sx -tx corresp_recov1 //; move: fx=>[cx _].
rewrite - xp.
move: (set_of_gfunctions_inc Vzg) => [x1 [fx1 [sx1 [tx1 gx1]]]].
suff: (x1 = (pr_i F z) \co x) by move => <-.
apply function_exten =>//; try fct_tac; try (rewrite /pr_i; aw; ue).
by apply pri_function =>//.
move=> x0 z0s.
have x0s: inc x0 src by ue.
have x0s': inc x0 (source x) by ue.
have cc: (pr_i F z) \coP x.
split; [apply pri_function=>//| split =>//; rewrite /pr_i; aw; ue].
have Wp: inc (W x0 x) (productb F) by rewrite -tx; apply inc_W_target.
aw; rewrite pri_W // /x; aw.
bw; rewrite/W; ue.
move=> t ts.
aw; split;gprops; bw; split=>//.
move=> i id; bw.
rewrite -dy in id; move: (py _ id); rewrite dy in id; bw => aux.
move: (set_of_gfunctions_inc aux) => [x2 [fx2 [sx2 [tx2 gx2]]]].
rewrite -tx2 -gx2; change (inc (W t x2) (target x2)).
by apply inc_W_target=>//; ue.
Qed.
End Bproduct.
Export Bproduct.
BL(fun f =>
L(domain F)( fun i=> (graph ( (pr_i F i) \co
(corresp src (productb F) f)))))
(set_of_gfunctions src (productb F))
(productb (L (domain F) (fun i=> set_of_gfunctions src (V i F)))).
Lemma fun_set_to_prod2: forall src F f gf,
fgraph F -> inc gf (set_of_gfunctions src (productb F)) ->
f = (corresp src (productb F) gf) -> function_prop f src (productb F).
Proof.
move=> src F f gf fF igf eq.
move: (set_of_gfunctions_inc igf) => [g [fg [sg [tg gg]]]].
have : f = g.
by rewrite eq -sg -tg -gg; apply corresp_recov1; move: fg=> [ res _].
move=> ->; red; intuition.
Qed.
Lemma fun_set_to_prod3: forall src F,
fgraph F -> transf_axioms(fun f =>
L(domain F)( fun i=> (graph (compose (pr_i F i)
(corresp src (productb F) f)))))
(set_of_gfunctions src (productb F))
(productb (L (domain F) (fun i=> set_of_gfunctions src (V i F)))).
Proof.
move=> src F fF x xs.
set (f:= corresp src (productb F) x).
move: (fun_set_to_prod2 fF xs (refl_equal f)) => [ff [sf tf]].
aw; gprops; split ;gprops; bw; split=>//.
move=> i id; bw.
move: (fun_set_to_prod1 fF id ff tf) => [fc [sc [tg vc]]].
by rewrite -sf -sc -tg; apply inc_set_of_gfunctions; gprops.
Qed.
Lemma fun_set_to_prod4: forall src F, fgraph F ->
function_prop (fun_set_to_prod src F) (set_of_gfunctions src (productb F))
(productb (L (domain F) (fun i=> set_of_gfunctions src (V i F)))).
Proof.
rewrite/fun_set_to_prod=>src F FF; red;aw;intuition.
by apply bl_function; apply fun_set_to_prod3.
Qed.
Definition fun_set_to_prod5 F f :=
ext_map_prod (domain F) (fun i=> source f)(fun i=> V i F)
(fun i => (graph (compose (pr_i F i) f))).
Lemma fun_set_to_prod6: forall F f,
fgraph F -> is_function f -> target f = productb F ->
(is_function (fun_set_to_prod5 F f) &
(fun_set_to_prod5 F f) \coP (constant_functor (domain F)(source f) )&
(fun_set_to_prod5 F f) \co (constant_functor (domain F)(source f)) =f).
Proof.
move=> F f fF ff tfpF.
set (tf := fun_set_to_prod5 F f).
have ftf: is_function tf.
rewrite /tf/fun_set_to_prod5; apply ext_map_prod_function.
move=> x xs.
move: (fun_set_to_prod1 fF xs ff tfpF) => [fc [sc [tc vc]]].
by split; fprops; aw; split=>//; rewrite -tc; apply f_range_graph.
have ttf:(productb F = target tf).
rewrite /tf/fun_set_to_prod5/ext_map_prod.
aw; rewrite /productf; apply f_equal; apply fgraph_exten; gprops; bw.
by move=> x xd; bw.
have ctf: tf \coP (constant_functor (domain F)(source f)).
split; fprops; split.
rewrite /constant_functor; apply bl_function.
move => x xs; rewrite graphset_pr2.
rewrite/cst_graph; bw; intuition; gprops.
by move=> y; bw; move=> [b [bf]] <-.
set (k:= cst_graph (domain F) (source f)).
have: (domain F = domain k) by rewrite/k /cst_graph; bw.
move=> ->; rewrite /constant_functor; aw.
rewrite - product_eq_graphset.
by rewrite /tf/fun_set_to_prod5/ext_map_prod; aw.
by rewrite /k /cst_graph; gprops.
by move => i; rewrite /k /cst_graph; rewrite L_domain=> id; bw.
split=> //; split=>//.
have ta1: transf_axioms [eta cst_graph (domain F)]
(source f) (set_of_gfunctions (domain F) (source f)).
move=> y yd; rewrite /constant_graph/set_of_gfunctions.
apply Z_inc; last by rewrite /cst_graph; bw; split; gprops.
apply powerset_inc; move=> t; rewrite /L fun_image_rw.
by move=> [z [zd]] <-; fprops.
symmetry; apply function_exten =>//; try fct_tac.
by rewrite /constant_functor; aw.
rewrite /tf/fun_set_to_prod5/ext_map_prod; aw.
by rewrite /productf tfpF L_recovers //.
move=> x xsf; rewrite /constant_functor; aw.
have cgd: inc (cst_graph (domain F) x)
(productf (domain F) (fun _ : Set => source f)).
rewrite/cst_graph; aw.
by split; gprops; bw; split; bw; move => j jd; bw.
have:inc (W x f) (productb F) by rewrite - tfpF; fprops.
have: (inc (W (cst_graph (domain F) x) tf) (productb F)).
rewrite ttf; apply inc_W_target=>//.
rewrite /tf/fun_set_to_prod5/ext_map_prod; aw.
rewrite /cst_graph.
by split;gprops; bw; split=>//; move => i id; bw.
move=> p1 p2; apply (productb_exten fF p2 p1).
move=> i idf; rewrite /tf/fun_set_to_prod5 ext_map_prod_W //.
bw; rewrite /ext_map_prod_aux /cst_graph; bw.
set cp:= (pr_i F i) \co f.
rewrite -[V x (graph cp)]/(W x cp) /cp;aw.
rewrite pri_W //.
split.
by apply pri_function=>//; split=>//.
by rewrite /pr_i; aw.
move => j jd.
move: (fun_set_to_prod1 fF jd ff tfpF) => [fcf [sc [tc vc]]].
split; fprops; split.
by rewrite -sc; fprops; aw.
by rewrite -tc;apply f_range_graph=>//.
Qed.
Lemma fun_set_to_prod7: forall src F f g ,
fgraph F ->
(forall i, inc i (domain F) -> inc (f i) (set_of_gfunctions src (V i F))) ->
g = ext_map_prod (domain F) (fun i=> src)(fun i=> V i F) f ->
(forall i, inc i (domain F) ->
f i = graph ((pr_i F i) \co (g \co (constant_functor (domain F) src) ))).
Proof.
move=> src F f g fF allf geq i idf.
set f1:= (constant_functor (domain F) src).
have fct: is_function f1.
rewrite /f1 /constant_functor; apply bl_function.
move=> x xd;rewrite graphset_pr2 /cst_graph.
by split; gprops ; bw; split =>//;move => w; bw; move=> [b [_] ] <-.
have fg: (is_function g).
rewrite geq; apply ext_map_prod_function.
move=> x xd; move:(set_of_gfunctions_inc (allf _ xd))=> [h [? [? [th gh]]]].
rewrite -gh f_domain_graph //; split; fprops; split=>//.
by rewrite -th; apply f_range_graph=>//.
have cpa: g \coP f1.
split=>//; split=>//.
rewrite geq/f1/ext_map_prod/constant_functor; aw.
rewrite /productf.
set (k:= (L (domain F) (fun _ : Set => src))).
have dkdf: (domain k = domain F) by rewrite /k;bw.
have fgk: (fgraph k) by rewrite /k; gprops.
have aux: (forall i, inc i(domain k) -> V i k = src).
by move=> j jd; rewrite /k; bw;ue.
by rewrite -dkdf; apply (product_eq_graphset fgk aux).
have fcf: is_function (g \co f1) by fct_tac.
have cpb: (pr_i F i) \coP (g \co f1).
split; first by apply pri_function => //.
split=>//; rewrite geq.
by rewrite /pr_i/ext_map_prod; aw; rewrite/productf L_recovers //.
have fctr: is_function ((pr_i F i) \co (g \co f1)) by fct_tac.
move: (allf _ idf) => aux; move: (set_of_gfunctions_inc aux).
move=> [h [fh [sh [th gh]]]].
have sf1: source f1 = source h by rewrite /f1/pr_i/constant_functor; aw.
have ta1: transf_axioms (fun x0 : Set => cst_graph (domain F) x0)
src (set_of_gfunctions (domain F) src).
move=> y yd; rewrite /cst_graph/set_of_gfunctions.
apply Z_inc; last by bw; split; gprops.
apply powerset_inc; move=> t; rewrite /L fun_image_rw.
by move=> [z [zd]] <-; fprops.
suff: (h = (pr_i F i) \co (g \co f1)) by move=> <-; ue.
apply function_exten =>//.
by rewrite /pr_i/constant_functor; aw.
by rewrite /pr_i/constant_functor; aw.
rewrite /f1; move => x xh.
have xs: inc x src by rewrite -sh.
have xs2: inc x (source (constant_functor (domain F) src)).
by rewrite /constant_functor; aw.
have aax: ext_map_prod_axioms (domain F) (fun _ : Set => src)
(fun i0 : Set => V i0 F) f.
move=> z zs; move: (allf _ zs) => p;move: (set_of_gfunctions_inc p).
move=> [k [fk [sk [tk gk]]]].
rewrite -gk -sk -tk f_domain_graph //; intuition.
by apply f_range_graph=>//.
set q:= cst_graph (domain F) x.
have qp: inc q (productf (domain F) (fun _ : Set => src)).
aw; rewrite/q /cst_graph; bw; split; gprops; split =>//.
by move=> k kd; bw.
have si: W x (constant_functor (domain F) src) = q.
by rewrite /W/constant_functor /BL; aw; bw.
have seq: (x = V i q) by rewrite constant_graph_V //.
aw; rewrite si geq ext_map_prod_W //.
set s:= (L (domain F) (ext_map_prod_aux q f)).
have isp: inc s (productb F).
rewrite /s /ext_map_prod_aux.
aw; split; gprops; bw; split =>//.
move=> j jd; bw; rewrite constant_graph_V //.
move: (allf _ jd) => p;move: (set_of_gfunctions_inc p).
by move=> [k [fk [sk [tk gk]]]]; rewrite -gk -tk; apply inc_W_target=>//; ue.
aw; rewrite pri_W => //.
by rewrite /s; bw; rewrite /ext_map_prod_aux /W gh -seq.
Qed.
Lemma fun_set_to_prod8: forall src F,
fgraph F -> bijection (fun_set_to_prod src F).
Proof.
move=> src F fF.
move: (fun_set_to_prod4 src fF).
set (g :=fun_set_to_prod src F) in *.
move => [fg [sg tg]].
split.
split =>//; move=> x y; rewrite sg => xs ys.
move :(set_of_gfunctions_inc xs) (set_of_gfunctions_inc ys).
move=> [x0 [fx0 [sx0 [tx0 gx]]]][x1 [fx1 [sx1 [tx1 gx1]]]] eq.
suff: x0 = x1 by rewrite - gx -gx1; move=>->.
move:(fun_set_to_prod3 (src:=src) fF)=> ta.
move: eq; rewrite /g/fun_set_to_prod; aw; clear ta; move => eq.
apply function_exten=>//; try ue.
move=> x2 x2s.
have iW0: (inc (W x2 x0) (productb F)) by ue.
have iW1: (inc (W x2 x1) (productb F)) by rewrite sx0 -sx1 in x2s; ue.
apply (productb_exten fF iW0 iW1).
move=> i iF.
move: (fun_set_to_prod1 fF iF fx0 tx0) (fun_set_to_prod1 fF iF fx1 tx1).
move=> [fc0 [sc0 [tc0 vc0]]] [fc1 [sc1 [tc1 vc1]]].
have x0c: (x0 = corresp src (productb F) x).
by rewrite -sx0 -tx0 -gx corresp_recov1 //; move: fx0=>[cx0 _].
have x1c: (x1 = corresp src (productb F) y).
by rewrite -sx1 -tx1 -gx1 corresp_recov1 //; move: fx1=>[cx1 _].
have cc: ((pr_i F i) \co x0 = (pr_i F i) \co x1).
move: (f_equal (V i) eq); bw; rewrite - x0c -x1c => eq1.
by apply function_exten1 =>//; aw.
by rewrite - vc0 //; rewrite sx0 -sx1 in x2s; rewrite cc - vc1 //.
apply surjective_pr6 => // y.
rewrite sg {1} /g /fun_set_to_prod bl_target; aw; gprops.
move=> [fy]; bw; move=> [dy py].
set (x:= BL (fun u=> L (domain F) (fun i=> V u (V i y))) src (productb F)).
have sx: source x = src by rewrite /x; aw.
have tx: target x = productb F by rewrite /x; aw.
have fx: is_function x.
rewrite /x; apply bl_function.
move=> t ts; aw; split; gprops; bw; split =>//.
move => i idf; bw.
rewrite dy in py; move: (py _ idf); bw=> xpp.
move: (set_of_gfunctions_inc xpp) => [k [fk [sk [tk gk]]]].
have: (V t (V i y) = W t k) by rewrite /W gk.
by move=>->; rewrite -tk; apply inc_W_target =>//; ue.
have gs: inc (graph x) (set_of_gfunctions src (productb F)).
rewrite -sx -tx; apply inc_set_of_gfunctions=> //.
ex_tac.
move: (fun_set_to_prod3 fF (src:=src)) => ax.
rewrite /g/fun_set_to_prod; aw; clear ax.
apply fgraph_exten; bw;gprops =>//.
move=> z zd; bw.
have zdd: inc z (domain y) by ue.
move: (py _ zdd); bw; move => Vzg.
have xp: (x = corresp src (productb F) (graph x)).
by rewrite -sx -tx corresp_recov1 //; move: fx=>[cx _].
rewrite - xp.
move: (set_of_gfunctions_inc Vzg) => [x1 [fx1 [sx1 [tx1 gx1]]]].
suff: (x1 = (pr_i F z) \co x) by move => <-.
apply function_exten =>//; try fct_tac; try (rewrite /pr_i; aw; ue).
by apply pri_function =>//.
move=> x0 z0s.
have x0s: inc x0 src by ue.
have x0s': inc x0 (source x) by ue.
have cc: (pr_i F z) \coP x.
split; [apply pri_function=>//| split =>//; rewrite /pr_i; aw; ue].
have Wp: inc (W x0 x) (productb F) by rewrite -tx; apply inc_W_target.
aw; rewrite pri_W // /x; aw.
bw; rewrite/W; ue.
move=> t ts.
aw; split;gprops; bw; split=>//.
move=> i id; bw.
rewrite -dy in id; move: (py _ id); rewrite dy in id; bw => aux.
move: (set_of_gfunctions_inc aux) => [x2 [fx2 [sx2 [tx2 gx2]]]].
rewrite -tx2 -gx2; change (inc (W t x2) (target x2)).
by apply inc_W_target=>//; ue.
Qed.
End Bproduct.
Export Bproduct.