Skip to content
Open
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
16 changes: 16 additions & 0 deletions CHANGELOG_UNRELEASED.md
Original file line number Diff line number Diff line change
Expand Up @@ -6,6 +6,22 @@
- in set_interval.v
+ definitions `itv_is_closed_unbounded`, `itv_is_cc`, `itv_closed_ends`

- in `tvs.v`
+ lemmas `cvg_sum`, `sum_continuous`

- in `unstable.v`:
+ structure `Norm`
+ lemmas `normMn`, `normN`, `ler_norm_sum`
+ definitions `max_norm`, `max_space`
+ lemmas `max_norm_ge0`, `le_coord_max_norm`, `max_norm0`, `ler_max_normD`,
`max_norm0_eq0`, `max_normZ`, `max_normMn`, `max_normN`

- in `normed_module.v`:
+ structure `NormedVector`
+ definition `normedVectType`
+ lemmas `sup_closed_ball_compact`, `equivalence_norms`,
`linear_findim_continuous`

### Changed
- in set_interval.v
+ `itv_is_open_unbounded`, `itv_is_oo`, `itv_open_ends` (Prop to bool)
Expand Down
114 changes: 113 additions & 1 deletion classical/unstable.v
Original file line number Diff line number Diff line change
@@ -1,5 +1,6 @@
(* mathcomp analysis (c) 2026 Inria and AIST. License: CeCILL-C. *)
From mathcomp Require Import all_ssreflect finmap ssralg ssrnum ssrint.
From HB Require Import structures.
From mathcomp Require Import all_ssreflect finmap ssralg ssrnum ssrint vector.
From mathcomp Require Import archimedean interval.

(**md**************************************************************************)
Expand Down Expand Up @@ -553,3 +554,114 @@ End ProperNotations.

Lemma sqrtK {K : rcfType} : {in Num.nneg, cancel (@Num.sqrt K) (fun x => x ^+ 2)}.
Proof. by move=> r r0; rewrite sqr_sqrtr. Qed.

Module Norm.

HB.mixin Record isNorm (K : numDomainType) (L : lmodType K) (norm : L -> K) := {
norm0 : norm 0 = 0;
norm_ge0 : forall x, 0 <= norm x;
norm0_eq0 : forall x, norm x = 0 -> x = 0;
ler_normD : forall x y, norm (x + y) <= norm x + norm y;
normZ : forall r x, norm (r *: x) = `|r| * norm x;
}.

#[export]
HB.structure Definition Norm K L := { norm of @isNorm K L norm }.

Module Import Theory.
Section Theory.
Variables (K : numDomainType) (L : lmodType K) (norm : Norm.type L).

Lemma normMn x n : norm (x *+ n) = norm x *+ n.
Proof. by rewrite -scaler_nat normZ normr_nat mulr_natl. Qed.

Lemma normN x : norm (- x) = norm x.
Proof. by rewrite -scaleN1r normZ normrN1 mul1r. Qed.

Lemma ler_norm_sum (I : Type) (r : seq I) (F : I -> L) :
norm (\sum_(i <- r) F i) <= \sum_(i <- r) norm (F i).
Proof.
by elim/big_ind2 : _ => *; rewrite ?norm0// (le_trans (ler_normD _ _))// lerD.
Qed.

End Theory.
End Theory.

Module Import Exports. HB.reexport. End Exports.
End Norm.
Export Norm.Exports.

Section InfiniteNorm.
Variables (K : numFieldType) (V : vectType K).
Let V' := @fullv _ V.
Variable B : (\dim V').-tuple V.
Hypothesis Bbasis : basis_of V' B.

Definition max_norm x := \big[Order.max/0]_(i < \dim V') `|coord B i x|.

Definition max_space : Type := (fun=> V) Bbasis.

HB.instance Definition _ := Vector.on max_space.

Lemma max_norm_ge0 x : 0 <= max_norm x.
Proof.
rewrite /max_norm.
by elim/big_ind : _ => //= ? ? ? ?; rewrite /Order.max; case: ifP.
Qed.

Lemma le_coord_max_norm x i : `|coord B i x| <= max_norm x.
Proof.
rewrite /max_norm; elim: (index_enum _) (mem_index_enum i) => //= j l IHl.
rewrite inE big_cons [X in _ <= X _ _]/Order.max/= => /predU1P[<-|/IHl {}IHl];
case: ifP => [/ltW|]// /negbT.
set b := (X in _ < X).
have bR : b \is Num.real by apply: bigmax_real => // a _; apply: normr_real.
have /comparable_leNgt <- := real_comparable bR (normr_real (coord B j x)).
by move=> /(le_trans IHl).
Qed.

Lemma max_norm0 : max_norm 0 = 0.
Proof.
apply: le_anti; rewrite max_norm_ge0 andbT.
apply: bigmax_le => // i _.
have <-: \sum_(i < \dim V') 0 *: B`_i = 0.
under eq_bigr do rewrite scale0r.
by rewrite sumr_const mul0rn.
by rewrite coord_sum_free ?normr0// (basis_free Bbasis).
Qed.

Lemma ler_max_normD x y : max_norm (x + y) <= max_norm x + max_norm y.
Proof.
apply: bigmax_le => [|/= i _]; first by rewrite addr_ge0// max_norm_ge0.
by rewrite raddfD/= (le_trans (ler_normD _ _))// lerD// le_coord_max_norm.
Qed.

Lemma max_norm0_eq0 x : max_norm x = 0 -> x = 0.
Proof.
move=> x0; rewrite (coord_basis Bbasis (memvf x)).
suff: forall i, coord B i x = 0.
by move=> {}x0; rewrite big1//= => j _; rewrite x0 scale0r.
by move=> i; apply/normr0_eq0/le_anti; rewrite normr_ge0 -x0 le_coord_max_norm.
Qed.

Lemma max_normZ r x : max_norm (r *: x) = `|r| * max_norm x.
Proof.
rewrite /max_norm.
under eq_bigr do rewrite linearZ/= normrM.
elim: (index_enum _) => [|i l IHl]; first by rewrite !big_nil mulr0.
by rewrite !big_cons IHl maxr_pMr.
Qed.

HB.instance Definition _ := Norm.isNorm.Build K V max_norm
max_norm0 max_norm_ge0 max_norm0_eq0 ler_max_normD max_normZ.

Lemma max_normMn x n : max_norm (x *+ n) = max_norm x *+ n.
Proof. exact: Norm.Theory.normMn. Qed.

Lemma max_normN x : max_norm (- x) = max_norm x.
Proof. exact: Norm.Theory.normN. Qed.

HB.instance Definition _ := Num.Zmodule_isNormed.Build K
max_space ler_max_normD max_norm0_eq0 max_normMn max_normN.

End InfiniteNorm.
2 changes: 1 addition & 1 deletion theories/normedtype_theory/complete_normed_module.v
Original file line number Diff line number Diff line change
@@ -1,6 +1,6 @@
(* mathcomp analysis (c) 2025 Inria and AIST. License: CeCILL-C. *)
From HB Require Import structures.
From mathcomp Require Import all_ssreflect ssralg ssrnum.
From mathcomp Require Import all_ssreflect ssralg ssrnum vector.
From mathcomp Require Import boolp classical_sets interval_inference reals.
From mathcomp Require Import topology tvs pseudometric_normed_Zmodule.
From mathcomp Require Import normed_module.
Expand Down
141 changes: 140 additions & 1 deletion theories/normedtype_theory/normed_module.v
Original file line number Diff line number Diff line change
Expand Up @@ -2,7 +2,7 @@
From HB Require Import structures.
From mathcomp Require Import all_ssreflect finmap ssralg ssrnum ssrint.
From mathcomp Require Import archimedean rat interval zmodp vector.
From mathcomp Require Import fieldext falgebra.
From mathcomp Require Import fieldext falgebra mathcomp_extra.
#[warning="-warn-library-file-internal-analysis"]
From mathcomp Require Import unstable.
From mathcomp Require Import boolp classical_sets filter functions cardinality.
Expand All @@ -21,6 +21,9 @@ From mathcomp Require Import ereal_normedtype pseudometric_normed_Zmodule.
(* normedModType K == interface type for a normed module *)
(* structure over the numDomainType K *)
(* The HB class is NormedModule. *)
(* normedVectType K == interface type for a normed vectType *)
(* structure over the numDomainType K *)
(* The HB class is NormedVector. *)
(* `|x| == the norm of x (notation from ssrnum.v) *)
(* ``` *)
(* *)
Expand Down Expand Up @@ -147,6 +150,10 @@ HB.instance Definition _ :=

HB.end.

#[short(type="normedVectType")]
HB.structure Definition NormedVector (K : numDomainType) :=
{T of NormedModule K T & Vector K T}.

(**md see also `Section standard_topology_pseudoMetricNormedZmod` in
`pseudometric_normed_Zmodule.v` *)
Section standard_topology_normedMod.
Expand Down Expand Up @@ -374,6 +381,9 @@ Unshelve. all: by end_near. Qed.
Lemma ball_open_nbhs (x : V) (r : K) : 0 < r -> open_nbhs x (ball x r).
Proof. by move=> e0; split; [exact: ball_open|exact: ballxx]. Qed.

HB.instance Definition _ := Norm.isNorm.Build K V (@Num.norm K V) (normr0 V)
(@normr_ge0 _ V) (@normr0_eq0 _ V) (@ler_normD _ V) (@normrZ _ V).

End NormedModule_numDomainType.

Definition self_sub (K : numDomainType) (V W : normedModType K)
Expand Down Expand Up @@ -2463,3 +2473,132 @@ Lemma open_disjoint_itv_bigcup : U = \bigcup_q open_disjoint_itv q.
Proof. by rewrite /open_disjoint_itv; case: cid => //= I [_]. Qed.

End open_set_disjoint_real_intervals.

Section InfiniteNorm.
Variables (R : numFieldType) (V : vectType R).
Let V' := @fullv _ V.
Variable B : (\dim V').-tuple V.
Hypothesis Bbasis : basis_of V' B.

HB.instance Definition _ := Pointed.copy (max_space Bbasis) V^o.

HB.instance Definition _ :=
PseudoMetric.copy (max_space Bbasis) (pseudoMetric_normed (max_space Bbasis)).

HB.instance Definition _ := NormedZmod_PseudoMetric_eq.Build R (max_space Bbasis) erefl.

HB.instance Definition _ :=
PseudoMetricNormedZmod_Lmodule_isNormedModule.Build R (max_space Bbasis) (max_normZ B).

End InfiniteNorm.

Section EquivalenceNorms.
Variables (R : realType) (V : vectType R).
Let V' := @fullv _ V.
Let Voo := max_space (vbasisP (@fullv _ V)).

(* N.B. Get Trocq to prove the continuity part automatically. *)
Lemma sup_closed_ball_compact : compact (closed_ball (0 : Voo) 1).
Proof.
rewrite closed_ballE// /closed_ball_.
under eq_set do rewrite sub0r normrN.
rewrite -[forall x, _]/(compact _).
pose f (X : {ptws 'I_(\dim V') -> R}) : Voo :=
\sum_(i < \dim V') X i *: (vbasis V')`_i.
have -> : [set x : Voo | `|x| <= 1] =
f @` [set X | forall i, `[-1, 1]%classic (X i)].
apply/seteqP; split=> [x/= x1|x/= [X X1 <-]].
- exists (coord (vbasis V') ^~ x); last first.
exact/esym/(@coord_vbasis _ _ (x : V))/memvf.
by move=> i; rewrite in_itv/= -ler_norml (le_trans _ x1)// le_coord_max_norm.
- rewrite /normr/= /max_norm bigmax_le => //= i _.
by rewrite coord_sum_free ?ler_norml; [exact: X1|exact/basis_free/vbasisP].
apply: (@continuous_compact _ _ f).
- apply/continuous_subspaceT/sum_continuous => /= i _ x.
exact/continuousZr_tmp/proj_continuous.
- apply: (@tychonoff 'I_(\dim V') (fun=> R^o) (fun=> `[-1, 1]%classic)) => _.
exact: segment_compact.
Qed.

Lemma equivalence_norms (N : Norm.Norm.type V) :
exists2 M, 0 < M & forall x : Voo, `|x| <= M * N x /\ N x <= M * `|x|.
Proof.
set M0 := 1 + \sum_(i < \dim V') N (vbasis V')`_i.
have M00 : 0 < M0 by rewrite ltr_pwDl// sumr_ge0// => ? _; apply: Norm.norm_ge0.
have leNoo (x : Voo) : N x <= M0 * `|x|.
rewrite [in leLHS](coord_vbasis (memvf (x : V))).
rewrite (le_trans (Norm.Theory.ler_norm_sum _ _ _))//.
rewrite mulrDl mul1r mulr_suml ler_wpDl// ler_sum => //= i _.
by rewrite Norm.normZ mulrC ler_pM// ?le_coord_max_norm// Norm.norm_ge0.
have NC0 : continuous (N : Voo -> R).
move=> /= x; rewrite /continuous_at.
apply: cvg_zero; first exact: nbhs_filter.
apply/cvgr0Pnorm_le; first exact: nbhs_filter.
move=> /= e e0.
near=> y.
rewrite -[_ y]/(N y - N x) (@le_trans _ _ (N (y - x)))//.
apply/ler_normlP.
have NB a b : N a <= N b + N (a - b).
by rewrite (le_trans _ (Norm.ler_normD _ _))// subrKC.
by rewrite opprB !lerBlDl NB -opprB Norm.Theory.normN NB.
rewrite (le_trans (leNoo _))// mulrC -ler_pdivlMr// -opprB normrN.
by near: y; apply: cvgr_dist_le; [exact: cvg_id|exact: divr_gt0].
have: compact [set x : Voo | `|x| = 1].
apply: (subclosed_compact _ sup_closed_ball_compact).
- apply: (@closed_comp _ _ _ [set 1 : R]); last exact: closed_eq.
by move=> *; exact: norm_continuous.
- by move => x/=; rewrite closed_ballE// /closed_ball_/= sub0r normrN => ->.
move=> /(@continuous_compact _ _ (N : Voo -> R)) -/(_ _)/wrap[].
exact: continuous_subspaceT.
move=> /(@continuous_compact _ _ (@GRing.inv R)) -/(_ _)/wrap[].
move=> /= x; rewrite /continuous_at.
apply: (@continuous_in_subspaceT _ _
[set N x | x in [set x : Voo | `|x| = 1]] (@GRing.inv R)).
move=> /= r /set_mem/= [y y1 <-].
apply: inv_continuous.
apply: contra_eq_neq y1 => /Norm.norm0_eq0 ->.
by rewrite normr0 eq_sym oner_eq0.
move=> /compact_bounded[M1 [M1R /(_ (1 + M1))]] /(_ (ltr_pwDl ltr01 (lexx _))).
rewrite /globally/= => M1N.
exists (maxr M0 (1 + M1)) => [|x]; first by rewrite lt_max M00.
split; last by rewrite (le_trans (leNoo x))// ler_wpM2r// le_max lexx.
have [->|x0] := eqVneq x 0; first by rewrite normr0 Norm.norm0 mulr0.
have Nx0 : 0 < N x.
rewrite lt0r Norm.norm_ge0 andbT.
by move: x0; apply: contra_neq => /Norm.norm0_eq0.
have normx0 : 0 < `|x| by rewrite normr_gt0.
move: M1N => /(_ (`|x| / N x)) -/(_ _)/wrap[].
exists (N x / `|x|); last by rewrite invf_div.
exists (`|x|^-1 *: x); last by rewrite Norm.normZ mulrC ger0_norm.
by rewrite normrZ normfV normr_id mulVf// gt_eqF.
rewrite ger0_norm; last by rewrite divr_ge0// Norm.norm_ge0.
rewrite ler_pdivrMr// => /le_trans; apply.
by rewrite ler_pM2r// le_max lexx orbT.
Unshelve. all: by end_near. Qed.

End EquivalenceNorms.

Section LinearContinuous.
Variables (R : realType) (V : normedVectType R) (W : normedModType R).

Lemma linear_findim_continuous (f : {linear V -> W}) : continuous f.
Proof.
set V' := @fullv _ V.
set B := vbasis V'.
move=> /= x; rewrite /continuous_at.
rewrite [x in f x](coord_vbasis (memvf x)) raddf_sum.
rewrite (@eq_cvg _ _ _ _ (fun y => \sum_(i < \dim V') coord B i y *: f B`_i));
last first.
move=> y; rewrite [y in LHS](coord_vbasis (memvf y)) raddf_sum.
by apply: eq_big => // i _; apply: linearZ.
apply: cvg_sum => i _.
rewrite [X in _ --> X]linearZ/= -/B.
apply: cvgZr_tmp.
move: x; apply/linear_bounded_continuous/bounded_funP => r/=.
have [M M0 MP] := equivalence_norms (@Num.norm _ V).
exists (M * r) => x.
move: MP => /(_ x) [Mx _] xr.
by rewrite (le_trans (le_coord_max_norm _ _ _))// (le_trans Mx) ?ler_wpM2l// ltW.
Qed.

End LinearContinuous.
16 changes: 16 additions & 0 deletions theories/tvs.v
Original file line number Diff line number Diff line change
Expand Up @@ -88,6 +88,22 @@ HB.mixin Record PreTopologicalNmodule_isTopologicalNmodule M
HB.structure Definition TopologicalNmodule :=
{M of PreTopologicalNmodule M & PreTopologicalNmodule_isTopologicalNmodule M}.

Section TopologicalNmodule_Theory.

Lemma cvg_sum (T : Type) (U : TopologicalNmodule.type) (F : set_system T)
(I : Type) (r : seq I) (P : pred I) (Ff : I -> T -> U) (Fa : I -> U) :
Filter F -> (forall i, P i -> Ff i x @[x --> F] --> Fa i) ->
\sum_(i <- r | P i) Ff i x @[x --> F] --> \sum_(i <- r| P i) Fa i.
Proof. by move=> FF Ffa; apply: cvg_big => //; apply: add_continuous. Qed.

Lemma sum_continuous (T : topologicalType) (M : TopologicalNmodule.type)
(I : Type) (r : seq I) (P : pred I) (F : I -> T -> M) :
(forall i : I, P i -> continuous (F i)) ->
continuous (fun x1 : T => \sum_(i <- r | P i) F i x1).
Proof. by move=> FC0; apply: continuous_big => //; apply: add_continuous. Qed.

End TopologicalNmodule_Theory.

HB.structure Definition PreTopologicalZmodule :=
{M of Topological M & GRing.Zmodule M}.

Expand Down
Loading