Skip to content

Commit 3388bf1

Browse files
committed
wip prod tvs
1 parent 5f7f946 commit 3388bf1

File tree

1 file changed

+36
-17
lines changed

1 file changed

+36
-17
lines changed

theories/tvs.v

Lines changed: 36 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -42,6 +42,18 @@ Definition convex (R : numDomainType) (M : lmodType R) (A : set M) :=
4242
forall x y (lambda : R), x \in A -> y \in A ->
4343
(0< lambda) -> (lambda < 1) -> lambda *: x + (1 - lambda) *: y \in A.
4444

45+
46+
Definition convex_hull (R : numDomainType) (M : lmodType R) (A : set M) :=
47+
[set z | exists2 l : R,(0 < l /\ (l < 1))&(exists2 x, (x \in A) &
48+
(exists2 y, (y \in A)&(z = (l *: x + (1 - l) *: y))))].
49+
50+
Lemma convex_convexhull (R : numDomainType) (M : lmodType R) (A : set M) :
51+
convex (convex_hull A).
52+
Proof.
53+
move => x y l; rewrite !inE /convex_hull /= =>- [lx [lx0 lx1] [x1 x1a] [x2 x2a]] xc.
54+
move => >- [ly [ly0 ly1] [y1 y1a] [y2 y2a]] yc l0 l1.
55+
Admitted.
56+
4557
HB.mixin Record Uniform_isTvs (R : numDomainType) E of Uniform E & GRing.Lmodule R E := {
4658
add_continuous : continuous (fun x : E * E => x.1 + x.2) ;
4759
(*continuous (uncurry (@GRing.add E))*)
@@ -327,23 +339,30 @@ Qed.
327339
Lemma prod_locally_convex :
328340
exists2 B : set (set (E * F)), (forall b, b \in B -> convex b) & basis B.
329341
Proof.
330-
move: (@locally_convex K E)=> [Be Bec Beb].
331-
move: (@locally_convex K F)=> [Bf Bfc Bfb].
332-
exists [set ef : set (E*F) |
333-
(exists be, exists2 bf, (be \in Be) & ((bf \in Bf) /\ (forall xy, (xy \in ef)<->(exists x, exists y, (x \in be) /\ (y \in bf) /\ (x,y)= xy))))].
334-
move=> b; rewrite inE /= => [[]] be [] bf Bee [] Bff H.
335-
rewrite /convex /= => ef1 ef2 l /H [x1 [y1]] [x1be] [y1bf] <- / H [x2 [y2]] [x2be] [y2bf] <- l0 l1.
336-
apply/H; exists (l*:x1 + (1-l)*:x2);exists (l*:y1 + (1-l)*:y2); split; first by apply: Bec.
337-
by split; first by apply: Bfc.
338-
split. move=> b /= => [[]] be [] bf ; rewrite !inE => Bee [] Bff H.
339-
suff: ((open be)/\(open bf)). admit.
340-
split; first by move: Beb => [] /= + _; apply.
341-
by move: Bfb => [] /= + _; apply.
342-
move => /= [x y] ef [[e f] /= [ne nf]] /= EF.
343-
exists [set z | (e z.1) /\ (f z.2)]; last by apply: EF.
344-
split; last by split; apply: nbhs_singleton.
345-
move: Beb=> [] _ /(_ x) /=.
346-
Admitted.
342+
move: (@locally_convex K E)=> [Be Bcb Beb].
343+
move: (@locally_convex K F)=> [Bf Bcf Bfb].
344+
pose B:= [set ef : set (E*F) | open ef /\
345+
(exists be, exists2 bf, (Be be) & (( Bf bf)/\ (be `*` bf = ef)))].
346+
exists B.
347+
move=> b; rewrite inE /= => [[]] bo [] be [] bf Bee [] Bff <-.
348+
move => [x1 y1] [x2 y2] l; rewrite !inE =>- /= [xe1 yf1] [xe2 yf2] l0 l1.
349+
split; rewrite -inE; first by apply: Bcb; rewrite ?inE.
350+
by apply: Bcf; rewrite ?inE.
351+
rewrite /basis /=.
352+
split; first by move=> b /= => [] [].
353+
move => /= [x y]; rewrite /filter_from /nbhs_simpl => ef [[ne nf]] /= [Ne Nf] Nef.
354+
rewrite nbhs_simpl /=.
355+
move: Beb=> [] Beo /(_ x ne Ne) /=; rewrite !nbhs_simpl /= =>- [a] [] Bea ax ea.
356+
move: Bfb=> [] Bfo /(_ y nf Nf) /=; rewrite !nbhs_simpl /= =>- [b] [] Beb yb fb.
357+
exists [set z | (a z.1) /\ (b z.2)]; last first.
358+
apply: subset_trans; last by apply:Nef.
359+
by move=> [zx zy] /= [] /ea + /fb.
360+
split => /=; last by split; rewrite /B /=.
361+
split; last by exists a; exists b; rewrite ?inE //.
362+
rewrite openE => [[z z'] /= [az bz]]; exists (a,b) => /=; last by [].
363+
rewrite !nbhsE /=; split; first by exists a => //; split => //; apply: Beo.
364+
by exists b => //; split => // []; apply: Bfo.
365+
Qed.
347366

348367
HB.instance Definition _ :=
349368
Uniform_isTvs.Build K (E * F)%type

0 commit comments

Comments
 (0)