|
| 1 | +(* mathcomp analysis (c) 2025 Inria and AIST. License: CeCILL-C. *) |
| 2 | +From HB Require Import structures. |
| 3 | +From mathcomp Require Import all_ssreflect ssralg ssrint ssrnum finmap matrix. |
| 4 | +From mathcomp Require Import rat interval zmodp vector fieldext falgebra. |
| 5 | +From mathcomp Require Import archimedean. |
| 6 | +From mathcomp Require Import mathcomp_extra unstable boolp classical_sets. |
| 7 | +From mathcomp Require Import functions cardinality set_interval. |
| 8 | +From mathcomp Require Import interval_inference ereal reals topology. |
| 9 | +From mathcomp Require Import function_spaces real_interval. |
| 10 | + |
| 11 | +Set Implicit Arguments. |
| 12 | +Unset Strict Implicit. |
| 13 | +Unset Printing Implicit Defensive. |
| 14 | + |
| 15 | +Import Order.TTheory GRing.Theory Num.Def Num.Theory. |
| 16 | +Import numFieldTopology.Exports. |
| 17 | + |
| 18 | +Local Open Scope classical_set_scope. |
| 19 | +Local Open Scope ring_scope. |
| 20 | + |
| 21 | +HB.mixin Record PseudoMetric_isMetric (K : realFieldType) M of PseudoMetric K M := { |
| 22 | + mdist : M -> M -> K ; |
| 23 | + mdist_ge0 : forall x y, 0 <= mdist x y ; |
| 24 | + mdist_positivity : forall x y, mdist x y = 0 -> x = y; |
| 25 | + ballEmdist : forall x d, ball x d = [set y | mdist x y < d] |
| 26 | +}. |
| 27 | + |
| 28 | +#[short(type="metricType")] |
| 29 | +HB.structure Definition MetricType (K : realFieldType) := |
| 30 | + { M of PseudoMetric K M & PseudoMetric_isMetric K M }. |
| 31 | + |
| 32 | +Section metric_lemmas. |
| 33 | +Variables (K : realFieldType) (M : metricType K). |
| 34 | + |
| 35 | +Let dist := @mdist K M. |
| 36 | + |
| 37 | +Lemma metric_sym x y : dist x y = dist y x. |
| 38 | +Proof. |
| 39 | +apply/eqP; rewrite eq_le; apply/andP; split; rewrite leNgt; apply/negP => xy. |
| 40 | +- have := @ball_sym _ _ y x (dist x y); rewrite !ballEmdist/= => /(_ xy). |
| 41 | + by rewrite ltxx. |
| 42 | +- have := @ball_sym _ _ x y (dist y x); rewrite !ballEmdist/= => /(_ xy). |
| 43 | + by rewrite ltxx. |
| 44 | +Qed. |
| 45 | + |
| 46 | +Lemma metricxx x : dist x x = 0. |
| 47 | +Proof. |
| 48 | +apply/eqP; rewrite eq_le mdist_ge0 andbT; apply/ler_addgt0Pl => /= e e0. |
| 49 | +rewrite addr0 leNgt; apply/negP => exx. |
| 50 | +by have := @ball_center _ _ x (PosNum e0); rewrite ballEmdist/= ltNge (ltW exx). |
| 51 | +Qed. |
| 52 | + |
| 53 | +Lemma metric_triangle x y z : dist x z <= dist x y + dist y z. |
| 54 | +Proof. |
| 55 | +apply/ler_addgt0Pl => /= e e0; rewrite leNgt; apply/negP => xyz. |
| 56 | +have := @ball_triangle _ _ y x z (dist x y + e / 2) (dist y z + e / 2). |
| 57 | +rewrite !ballEmdist/= addrACA -splitr !ltrDl !divr_gt0// => /(_ isT isT). |
| 58 | +by rewrite ltNge => /negP; apply; rewrite addrC ltW. |
| 59 | +Qed. |
| 60 | + |
| 61 | +Lemma metric_hausdorff : hausdorff_space M. |
| 62 | +Proof. |
| 63 | +move=> p q pq; apply: contrapT => qp; move: pq. |
| 64 | +pose D := dist p q / 2; have D0 : 0 < D. |
| 65 | + rewrite divr_gt0// lt_neqAle mdist_ge0 andbT eq_sym. |
| 66 | + by move/eqP: qp; apply: contra => /eqP/mdist_positivity ->. |
| 67 | +have p2Dq : ~ ball p (dist p q) q by rewrite ballEmdist/= ltxx. |
| 68 | +move=> /(_ (ball p _) (ball q _) (nbhsx_ballx _ _ D0) (nbhsx_ballx _ _ D0)). |
| 69 | +move/set0P/eqP; apply; rewrite -subset0 => x [pDx /ball_sym qDx]. |
| 70 | +by have := ball_triangle pDx qDx; rewrite -splitr. |
| 71 | +Qed. |
| 72 | + |
| 73 | +End metric_lemmas. |
0 commit comments