Skip to content

Commit f705646

Browse files
committed
[B] Simplify def. of continuous_2arg
This hides the destructuring of the pair inside the projection maps. This makes proofs of `continuous_2arg` a lot easier, because now the continuity of the projections and continuity of composition can be used easier than before.
1 parent f01100a commit f705646

File tree

2 files changed

+9
-9
lines changed

2 files changed

+9
-9
lines changed

theories/Topology/ProductTopology.v

Lines changed: 4 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -358,11 +358,11 @@ Variable f:point_set X -> point_set Y -> point_set Z.
358358

359359
Definition continuous_2arg :=
360360
continuous (fun p:point_set X * point_set Y =>
361-
let (x,y):=p in f x y)
361+
f (fst p) (snd p))
362362
(X:=ProductTopology2 X Y).
363363
Definition continuous_at_2arg (x:point_set X) (y:point_set Y) :=
364364
continuous_at (fun p:point_set X * point_set Y =>
365-
let (x,y):=p in f x y) (x, y)
365+
f (fst p) (snd p)) (x, y)
366366
(X:=ProductTopology2 X Y).
367367

368368
Lemma continuous_2arg_func_continuous_everywhere:
@@ -398,9 +398,10 @@ Lemma continuous_composition_at_2arg:
398398
continuous_at (fun w:point_set W => f (g w) (h w)) w.
399399
Proof.
400400
intros.
401+
red in H.
401402
apply (continuous_composition_at
402403
(fun p:point_set (ProductTopology2 X Y) =>
403-
let (x,y):=p in f x y)
404+
f (fst p) (snd p))
404405
(fun w:point_set W => (g w, h w))); trivial.
405406
now apply product2_map_continuous_at.
406407
Qed.

theories/Topology/RFuncContinuity.v

Lines changed: 5 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -81,6 +81,7 @@ apply continuous_at_neighborhood_basis with
8181
+ destruct x0 as [x' y'],
8282
H1 as [[[] []]].
8383
unfold R_metric.
84+
simpl.
8485
replace (x'+y' - (x+y)) with ((x'-x) + (y'-y)) by ring.
8586
apply Rle_lt_trans with (Rabs (x'-x) + Rabs(y'-y)).
8687
* apply Rabs_triang.
@@ -167,9 +168,8 @@ pose proof (RTop_metrization 0).
167168
apply continuous_at_neighborhood_basis with
168169
(metric_topology_neighborhood_basis R_metric 0).
169170
- apply open_neighborhood_basis_is_neighborhood_basis.
170-
replace (0*0) with 0 by auto with real.
171-
apply H.
172-
171+
simpl. rewrite Rmult_0_r.
172+
assumption.
173173
- intros.
174174
destruct H0.
175175
exists (characteristic_function_to_ensemble
@@ -192,9 +192,8 @@ apply continuous_at_neighborhood_basis with
192192
destruct x as [x y].
193193
destruct H1 as [[] []].
194194
unfold R_metric in *.
195-
replace (x*y-0) with (x*y) by auto with real.
196-
replace (x-0) with x in H1 by auto with real.
197-
replace (y-0) with y in H2 by auto with real.
195+
simpl.
196+
rewrite Rminus_0_r in *.
198197
rewrite Rabs_mult.
199198
replace r with (r*1) by auto with real.
200199
apply Rmult_le_0_lt_compat;

0 commit comments

Comments
 (0)