Skip to content

Commit 38a49d4

Browse files
committed
instances for dependent function types
1 parent d26b17b commit 38a49d4

File tree

11 files changed

+165
-74
lines changed

11 files changed

+165
-74
lines changed

altreals/realseq.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -303,7 +303,7 @@ Lemma ncvgM u v lu lv : ncvg u lu%:E -> ncvg v lv%:E ->
303303
Proof.
304304
move=> cu cv; pose a := u \- lu%:S; pose b := v \- lv%:S.
305305
have eq: (u \* v) =1 (lu * lv)%:S \+ ((lu%:S \* b) \+ (a \* v)).
306-
move=> n; rewrite {}/a {}/b /= [u n+_]addrC [(_+_)*(v n)]mulrDl.
306+
move=> n; rewrite {}/a {}/b /= [u n+_]addrC [(_+_)*(v n)]mulrDl.
307307
rewrite !addrA -[LHS]add0r; congr (_ + _); rewrite mulrDr.
308308
by rewrite !(mulrN, mulNr) (addrCA (lu * lv)) subrr addr0 subrr.
309309
apply/(ncvg_eq eq); rewrite -[X in X%:E]addr0; apply/ncvgD.

classical/cardinality.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1340,7 +1340,7 @@ Section zmod.
13401340
Context (aT : Type) (rT : zmodType).
13411341
Lemma fimfun_zmod_closed : zmod_closed (@fimfun aT rT).
13421342
Proof.
1343-
split=> [|f g]; rewrite !inE/=; first exact: finite_image_cst.
1343+
split=> [|f g]; rewrite !inE/=; first exact: (finite_image_cst 0).
13441344
by move=> fA gA; apply: (finite_image11 (fun x y => x - y)).
13451345
Qed.
13461346
HB.instance Definition _ :=

classical/functions.v

Lines changed: 102 additions & 30 deletions
Original file line numberDiff line numberDiff line change
@@ -2568,61 +2568,133 @@ Qed.
25682568

25692569
Obligation Tactic := idtac.
25702570

2571-
Program Definition fct_zmodMixin (T : Type) (M : zmodType) :=
2572-
@GRing.isZmodule.Build (T -> M) \0 (fun f x => - f x) (fun f g => f \+ g)
2571+
Definition add_fun {aT} {rT : aT -> nmodType} (f g : forall x, rT x) x := f x + g x.
2572+
Definition opp_fun {aT} {rT : aT -> zmodType} (f : forall x, rT x) x := - f x.
2573+
Definition sub_fun {aT} {rT : aT -> zmodType} (f g : forall x, rT x) x := f x - g x.
2574+
Definition mul_fun {aT} {rT : aT -> semiRingType} (f g : forall x, rT x) x := f x * g x.
2575+
Definition scale_fun {aT} {R} {rT : aT -> lmodType R} (k : R) (f : forall x, rT x) x := k *: f x.
2576+
2577+
Notation "f \+ g" := (add_fun f g) : ring_scope.
2578+
Notation "\- f" := (opp_fun f) : ring_scope.
2579+
Notation "f \- g" := (sub_fun f g) : ring_scope.
2580+
Notation "f \* g" := (mul_fun f g) : ring_scope.
2581+
Notation "k \*: f" := (scale_fun k f) : ring_scope.
2582+
2583+
Arguments add_fun {_ _} f g _ /.
2584+
Arguments opp_fun {_ _} f _ /.
2585+
Arguments sub_fun {_ _} f g _ /.
2586+
Arguments mul_fun {_ _} f g _ /.
2587+
Arguments scale_fun {_ _ _} k f _ /.
2588+
2589+
Program Definition fct_zmodMixin (T : Type) (M : T -> zmodType) :=
2590+
@GRing.isZmodule.Build (forall t, M t) (fun=> 0) (fun f => \- f) (fun f g => f \+ g)
25732591
_ _ _ _.
2574-
Next Obligation. by move=> T M f g h; rewrite funeqE=> x /=; rewrite addrA. Qed.
2575-
Next Obligation. by move=> T M f g; rewrite funeqE=> x /=; rewrite addrC. Qed.
2576-
Next Obligation. by move=> T M f; rewrite funeqE=> x /=; rewrite add0r. Qed.
2577-
Next Obligation. by move=> T M f; rewrite funeqE=> x /=; rewrite addNr. Qed.
2578-
HB.instance Definition _ (T : Type) (M : zmodType) := fct_zmodMixin T M.
2579-
2580-
Program Definition fct_ringMixin (T : pointedType) (M : ringType) :=
2581-
@GRing.Zmodule_isRing.Build (T -> M) (cst 1) (fun f g => f \* g) _ _ _ _ _ _.
2582-
Next Obligation. by move=> T M f g h; rewrite funeqE=> x /=; rewrite mulrA. Qed.
2583-
Next Obligation. by move=> T M f; rewrite funeqE=> x /=; rewrite mul1r. Qed.
2584-
Next Obligation. by move=> T M f; rewrite funeqE=> x /=; rewrite mulr1. Qed.
2585-
Next Obligation. by move=> T M f g h; rewrite funeqE=> x/=; rewrite mulrDl. Qed.
2586-
Next Obligation. by move=> T M f g h; rewrite funeqE=> x/=; rewrite mulrDr. Qed.
25872592
Next Obligation.
2588-
by move=> T M ; apply/eqP; rewrite funeqE => /(_ point) /eqP; rewrite oner_eq0.
2593+
by move=> T M f g h; apply/functional_extensionality_dep => x /=; rewrite addrA.
25892594
Qed.
2590-
HB.instance Definition _ (T : pointedType) (M : ringType) := fct_ringMixin T M.
2595+
Next Obligation.
2596+
by move=> T M f g; apply/functional_extensionality_dep => x /=; rewrite addrC.
2597+
Qed.
2598+
Next Obligation.
2599+
by move=> T M f; apply/functional_extensionality_dep => x /=; rewrite add0r.
2600+
Qed.
2601+
Next Obligation.
2602+
by move=> T M f; apply/functional_extensionality_dep => x /=; rewrite addNr.
2603+
Qed.
2604+
HB.instance Definition _ (T : Type) (M : T -> zmodType) := fct_zmodMixin M.
2605+
2606+
Program Definition fct_ringMixin (T : pointedType) (M : T -> ringType) :=
2607+
@GRing.Zmodule_isRing.Build (forall t, M t) (fun=> 1) (fun f g => f \* g) _ _ _ _ _ _.
2608+
Next Obligation.
2609+
by move=> T M f g h; apply/functional_extensionality_dep => x /=; rewrite mulrA.
2610+
Qed.
2611+
Next Obligation.
2612+
by move=> T M f; apply/functional_extensionality_dep => x /=; rewrite mul1r.
2613+
Qed.
2614+
Next Obligation.
2615+
by move=> T M f; apply/functional_extensionality_dep => x /=; rewrite mulr1.
2616+
Qed.
2617+
Next Obligation.
2618+
by move=> T M f g h; apply/functional_extensionality_dep => x/=; rewrite mulrDl.
2619+
Qed.
2620+
Next Obligation.
2621+
by move=> T M f g h; apply/functional_extensionality_dep => x/=; rewrite mulrDr.
2622+
Qed.
2623+
Next Obligation.
2624+
by move=> T M ; apply/eqP => /(congr1 (fun f => f point)) /eqP; rewrite oner_eq0.
2625+
Qed.
2626+
HB.instance Definition _ (T : pointedType) (M : T -> ringType) := fct_ringMixin M.
25912627

2592-
Program Definition fct_comRingType (T : pointedType) (M : comRingType) :=
2593-
GRing.Ring_hasCommutativeMul.Build (T -> M) _.
2628+
Program Definition fct_comRingType (T : pointedType) (M : T -> comRingType) :=
2629+
GRing.Ring_hasCommutativeMul.Build (forall t, M t) _.
25942630
Next Obligation.
2595-
by move=> T M f g; rewrite funeqE => x; rewrite /GRing.mul/= mulrC.
2631+
by move=> T M f g; apply/functional_extensionality_dep => x; apply/mulrC.
25962632
Qed.
2597-
HB.instance Definition _ (T : pointedType) (M : comRingType) :=
2598-
fct_comRingType T M.
2633+
HB.instance Definition _ (T : pointedType) (M : T -> comRingType) :=
2634+
fct_comRingType M.
25992635

26002636
Section fct_lmod.
2601-
Variables (U : Type) (R : ringType) (V : lmodType R).
2602-
Program Definition fct_lmodMixin := @GRing.Zmodule_isLmodule.Build R (U -> V)
2637+
Variables (U : Type) (R : ringType) (V : U -> lmodType R).
2638+
Program Definition fct_lmodMixin := @GRing.Zmodule_isLmodule.Build R (forall u, V u)
26032639
(fun k f => k \*: f) _ _ _ _.
2604-
Next Obligation. by move=> k f v; rewrite funeqE=> x; exact: scalerA. Qed.
2605-
Next Obligation. by move=> f; rewrite funeqE=> x /=; rewrite scale1r. Qed.
26062640
Next Obligation.
2607-
by move=> f g h; rewrite funeqE => x /=; rewrite scalerDr.
2641+
by move=> k f v; apply/functional_extensionality_dep => x; exact: scalerA.
2642+
Qed.
2643+
Next Obligation.
2644+
by move=> f; apply/functional_extensionality_dep => x /=; rewrite scale1r.
2645+
Qed.
2646+
Next Obligation.
2647+
by move=> f g h; apply/functional_extensionality_dep => x /=; rewrite scalerDr.
26082648
Qed.
26092649
Next Obligation.
2610-
by move=> f g h; rewrite funeqE => x /=; rewrite scalerDl.
2650+
by move=> f g h; apply/functional_extensionality_dep => x /=; rewrite scalerDl.
26112651
Qed.
26122652
HB.instance Definition _ := fct_lmodMixin.
26132653
End fct_lmod.
26142654

2615-
Lemma fct_sumE (I T : Type) (M : zmodType) r (P : {pred I}) (f : I -> T -> M)
2655+
Lemma add_funE (T : pointedType) (M : T -> zmodType) (f g : forall t, M t) :
2656+
(f + g) = f \+ g.
2657+
Proof. exact/functional_extensionality_dep. Qed.
2658+
2659+
Lemma opp_funE (T : Type) (M : T -> zmodType) (f : forall t, M t) : - f = \- f.
2660+
Proof. exact/functional_extensionality_dep. Qed.
2661+
2662+
Lemma sub_funE (T : Type) (M : T -> zmodType) (f g : forall t, M t) :
2663+
(f - g) = f \- g.
2664+
Proof. exact/functional_extensionality_dep. Qed.
2665+
2666+
Lemma fct_sumE (I T : Type) (M : T -> zmodType) r (P : {pred I}) (f : I -> forall t, M t)
26162667
(x : T) :
26172668
(\sum_(i <- r | P i) f i) x = \sum_(i <- r | P i) f i x.
26182669
Proof. by elim/big_rec2: _ => //= i y ? Pi <-. Qed.
26192670

2671+
Lemma mul_funE (T : pointedType) (M : T -> ringType) (f g : forall t, M t) :
2672+
(f * g) = f \* g.
2673+
Proof. exact/functional_extensionality_dep. Qed.
2674+
2675+
Lemma scale_funE (T : Type) (R : ringType) (M : T -> lmodType R) (k : R) (f : forall t, M t) :
2676+
k *: f = k \*: f.
2677+
Proof. exact/functional_extensionality_dep. Qed.
2678+
26202679
Lemma mul_funC (T : Type) {R : comSemiRingType} (f : T -> R) (r : R) :
26212680
r \*o f = r \o* f.
2622-
Proof. by apply/funext => x/=; rewrite mulrC. Qed.
2681+
Proof.
2682+
by apply/funext => x/=; rewrite mulrC. Qed.
26232683

26242684
End function_space.
26252685

2686+
Notation "f \+ g" := (add_fun f g) : ring_scope.
2687+
Notation "\- f" := (opp_fun f) : ring_scope.
2688+
Notation "f \- g" := (sub_fun f g) : ring_scope.
2689+
Notation "f \* g" := (mul_fun f g) : ring_scope.
2690+
Notation "k \*: f" := (scale_fun k f) : ring_scope.
2691+
2692+
Arguments add_fun {_ _} f g _ /.
2693+
Arguments opp_fun {_ _} f _ /.
2694+
Arguments sub_fun {_ _} f g _ /.
2695+
Arguments mul_fun {_ _} f g _ /.
2696+
Arguments scale_fun {_ _ _} k f _ /.
2697+
26262698
Section function_space_lemmas.
26272699
Local Open Scope ring_scope.
26282700
Import GRing.Theory.

theories/derive.v

Lines changed: 26 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -224,7 +224,7 @@ Section diff_locally_converse_tentative.
224224
(* and thus a littleo of 1, and so is id *)
225225
(* This can be generalized to any dimension *)
226226
Lemma diff_locally_converse_part1 (f : R -> R) (a b x : R) :
227-
f \o shift x = cst a + b *: idfun +o_ 0 id -> f x = a.
227+
f \o shift x = cst a + b *: (@idfun R) +o_ 0 id -> f x = a.
228228
Proof.
229229
rewrite funeqE => /(_ 0) /=; rewrite add0r => ->.
230230
by rewrite -[LHS]/(_ 0 + _ 0 + _ 0) /cst [X in a + X]scaler0 littleo_lim0 ?addr0.
@@ -253,12 +253,12 @@ Lemma derivable_nbhs (f : V -> W) a v :
253253
Proof.
254254
move=> df; apply/eqaddoP => _/posnumP[e].
255255
rewrite -nbhs_nearE nbhs_simpl /= dnbhsE; split; last first.
256-
rewrite /at_point opprD -![(_ + _ : _ -> _) _]/(_ + _) scale0r add0r.
256+
rewrite /at_point opprD !add_funE !opp_funE/= scale0r add0r.
257257
by rewrite addrA subrr add0r normrN scale0r !normr0 mulr0.
258258
have /eqolimP := df.
259259
move=> /eqaddoP /(_ e%:num) /(_ [gt0 of e%:num]).
260260
apply: filter_app; rewrite /= !near_simpl near_withinE; near=> h => hN0.
261-
rewrite /= opprD -![(_ + _ : _ -> _) _]/(_ + _) -![(- _ : _ -> _) _]/(- _).
261+
rewrite /= opprD !add_funE !opp_funE.
262262
rewrite /cst /= [`|1|]normr1 mulr1 => dfv.
263263
rewrite addrA -[X in X + _]scale1r -(@mulVf _ h) //.
264264
rewrite mulrC -scalerA -scalerBr normrZ.
@@ -276,7 +276,7 @@ move=> df; apply/cvg_ex; exists ('D_v f a).
276276
apply/(@eqolimP _ _ _ (dnbhs_filter_on _))/eqaddoP => _/posnumP[e].
277277
have /eqaddoP /(_ e%:num) /(_ [gt0 of e%:num]) := df.
278278
rewrite /= !(near_simpl, near_withinE); apply: filter_app; near=> h.
279-
rewrite /= opprD -![(_ + _ : _ -> _) _]/(_ + _) -![(- _ : _ -> _) _]/(- _).
279+
rewrite /= opprD !add_funE !opp_funE.
280280
rewrite /cst /= [`|1|]normr1 mulr1 addrA => dfv hN0.
281281
rewrite -[X in _ - X]scale1r -(@mulVf _ h) //.
282282
rewrite -scalerA -scalerBr normrZ normfV ler_pdivrMl ?normr_gt0 //.
@@ -398,7 +398,7 @@ Variable R : numFieldType.
398398
Fact dcst (V W : normedModType R) (a : W) (x : V) : continuous (0 : V -> W) /\
399399
cst a \o shift x = cst (cst a x) + \0 +o_ 0 id.
400400
Proof.
401-
split; first exact: cst_continuous.
401+
split; first exact: (@cst_continuous _ _ 0).
402402
apply/eqaddoE; rewrite addr0 funeqE => ? /=; rewrite -[LHS]addr0; congr (_ + _).
403403
by rewrite littleoE; last exact: littleo0_subproof.
404404
Qed.
@@ -412,7 +412,7 @@ Fact dadd (f g : V -> W) x :
412412
Proof.
413413
move=> df dg; split => [?|]; do ?exact: continuousD.
414414
apply/(@eqaddoE R); rewrite funeqE => y /=; rewrite -[(f + g) _]/(_ + _).
415-
by rewrite ![_ (_ + x)]diff_locallyx// addrACA addox addrACA.
415+
by rewrite add_funE ![_ (_ + x)]diff_locallyx// addrACA addox addrACA.
416416
Qed.
417417

418418
Fact dopp (f : V -> W) x :
@@ -521,8 +521,8 @@ Lemma diff_unique (V W : normedModType R) (f : V -> W)
521521
continuous df -> f \o shift x = cst (f x) + df +o_ 0 id ->
522522
'd f x = df :> (V -> W).
523523
Proof.
524-
move=> dfc dxf; apply/subr0_eq; rewrite -[LHS]/(_ \- _).
525-
apply/littleo_linear0/eqoP/eq_some_oP => /=; rewrite funeqE => y /=.
524+
move=> dfc dxf; apply/subr0_eq/(littleo_linear0 (f:=GRing.sub_fun _ _))/eqoP.
525+
apply/eq_some_oP => /=; rewrite funeqE => y /=.
526526
have hdf h : (f \o shift x = cst (f x) + h +o_ 0 id) ->
527527
h = f \o shift x - cst (f x) +o_ 0 id.
528528
move=> hdf; apply: eqaddoE.
@@ -544,7 +544,7 @@ by rewrite littleo_center0 (comp_centerK x id) -[- _ in RHS](comp_centerK x).
544544
Qed.
545545

546546
Lemma diff_cst (V W : normedModType R) a x : ('d (cst a) x : V -> W) = 0.
547-
Proof. by apply/diff_unique; have [] := dcst a x. Qed.
547+
Proof. by apply/(diff_unique (df:=\0)); have [] := dcst a x. Qed.
548548

549549
Variables (V W : normedModType R).
550550

@@ -558,7 +558,10 @@ Proof. exact: DiffDef (differentiable_cst _ _) (diff_cst _ _). Qed.
558558
Lemma diffD (f g : V -> W) x :
559559
differentiable f x -> differentiable g x ->
560560
'd (f + g) x = 'd f x \+ 'd g x :> (V -> W).
561-
Proof. by move=> df dg; apply/diff_unique; have [] := dadd df dg. Qed.
561+
Proof.
562+
move=> df dg.
563+
by apply/(diff_unique (df:=GRing.add_fun _ _)); have [] := dadd df dg.
564+
Qed.
562565

563566
Lemma differentiableD (f g : V -> W) x :
564567
differentiable f x -> differentiable g x -> differentiable (f + g) x.
@@ -576,7 +579,8 @@ Qed.
576579
Lemma differentiable_sum n (f : 'I_n -> V -> W) (x : V) :
577580
(forall i, differentiable (f i) x) -> differentiable (\sum_(i < n) f i) x.
578581
Proof.
579-
by elim/big_ind : _ => // ? ? g h ?; apply: differentiableD; [exact:g|exact:h].
582+
elim/big_ind : _ => //[_|? ? g h ?]; first exact/(@differentiable_cst _ 0).
583+
apply: differentiableD; [exact:g|exact:h].
580584
Qed.
581585

582586
Lemma diffN (f : V -> W) x :
@@ -614,7 +618,10 @@ Proof. by move=> /differentiableP df /differentiableP dg. Qed.
614618

615619
Lemma diffZ (f : V -> W) k x :
616620
differentiable f x -> 'd (k *: f) x = k \*: 'd f x :> (V -> W).
617-
Proof. by move=> df; apply/diff_unique; have [] := dscale k df. Qed.
621+
Proof.
622+
move=> df.
623+
by apply/(diff_unique (df:=GRing.scale_fun _ _)); have [] := dscale k df.
624+
Qed.
618625

619626
Lemma differentiableZ (f : V -> W) k x :
620627
differentiable f x -> differentiable (k *: f) x.
@@ -1152,7 +1159,7 @@ Global Instance is_derive_sum n (h : 'I_n -> V -> W) (x v : V)
11521159
(dh : 'I_n -> W) : (forall i, is_derive x v (h i) (dh i)) ->
11531160
is_derive x v (\sum_(i < n) h i) (\sum_(i < n) dh i).
11541161
Proof.
1155-
by elim/big_ind2 : _ => // [|] *; [exact: is_derive_cst|exact: is_deriveD].
1162+
by elim/big_ind2 : _ => // [|] *; [exact/(is_derive_cst 0)|exact: is_deriveD].
11561163
Qed.
11571164

11581165
Lemma derivable_sum n (h : 'I_n -> V -> W) (x v : V) :
@@ -1324,7 +1331,10 @@ by rewrite [in LHS]mulr_natl exprfctE -mulrSr mulr_natl.
13241331
Qed.
13251332

13261333
Lemma derivableX f n (x v : V) : derivable f x v -> derivable (f ^+ n) x v.
1327-
Proof. by case: n => [_|n /derivableP]; [rewrite expr0|]. Qed.
1334+
Proof.
1335+
case: n => [_|n /derivableP]; last by [].
1336+
by rewrite expr0; apply/(derivable_cst (1 : R)).
1337+
Qed.
13281338

13291339
Lemma deriveX f n (x v : V) : derivable f x v ->
13301340
'D_v (f ^+ n.+1) x = (n.+1%:R * f x ^+ n) *: 'D_v f x.
@@ -1378,7 +1388,7 @@ Proof. by rewrite derive1E derive_cst. Qed.
13781388
Lemma exprn_derivable {R : numFieldType} n (x : R) v :
13791389
derivable (@GRing.exp R ^~ n) x v.
13801390
Proof.
1381-
elim: n => [/=|n ih]; first by rewrite (_ : _ ^~ _ = 1).
1391+
elim: n => [/=|n ih]; first by rewrite (_ : _ ^~ _ = cst 1).
13821392
rewrite (_ : _ ^~ _ = (fun x => x * x ^+ n)); last first.
13831393
by apply/funext => y; rewrite exprS.
13841394
by apply: derivableM; first exact: derivable_id.
@@ -1573,7 +1583,7 @@ have gcont : {within `[a, b], continuous g}.
15731583
move=> x; apply: continuousD _ ; first by move=>?; exact: fcont.
15741584
by apply/continuousN/continuous_subspaceT=> ?; exact: scalel_continuous.
15751585
have gaegb : g a = g b.
1576-
rewrite /g -![(_ - _ : _ -> _) _]/(_ - _).
1586+
rewrite /g !add_funE !opp_funE.
15771587
apply/eqP; rewrite -subr_eq /= opprK addrAC -addrA -scalerBl.
15781588
rewrite [_ *: _]mulrA mulrC mulrA mulVf.
15791589
by rewrite mul1r addrCA subrr addr0.

theories/esum.v

Lines changed: 3 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -587,9 +587,10 @@ move=> /cvgrPdist_lt cvgAB; apply/cvgrPdist_lt => e e0.
587587
move: cvgAB => /(_ _ e0) [N _/= hN] /=.
588588
near=> n.
589589
rewrite distrC subr0.
590-
have -> : (C_ = A_ \- B_)%R.
590+
have -> : (C_ = A_ - B_)%R.
591591
apply/funext => k.
592-
rewrite /= /A_ /C_ /B_ -sumrN -big_split/= -summable_fine_sum//.
592+
rewrite /= /A_ /C_ /B_ sub_funE/=.
593+
rewrite -sumrN -big_split/= -summable_fine_sum//.
593594
apply eq_bigr => i Pi; rewrite -fineB//.
594595
- by rewrite [in LHS](funeposneg f).
595596
- by rewrite fin_num_abs (@summable_pinfty _ _ P) //; exact/summable_funepos.

theories/ftc.v

Lines changed: 9 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -839,7 +839,8 @@ set f := fun x => if x == a then r else if x == b then l else F^`() x.
839839
have fE : {in `]a, b[, F^`() =1 f}.
840840
by move=> x; rewrite in_itv/= => /andP[ax xb]; rewrite /f gt_eqF// lt_eqF.
841841
have DPGFE : {in `]a, b[, (- (PG \o F))%R^`() =1 ((G \o F) * (- f))%R}.
842-
move=> x /[dup]xab /andP[ax xb]; rewrite derive1_comp //; last first.
842+
move=> x /[dup]xab /andP[ax xb].
843+
rewrite (derive1_comp (g:[email protected] R)) //; last first.
843844
apply: diff_derivable; apply: differentiable_comp; apply/derivable1_diffP.
844845
by case: Fab => + _ _; exact.
845846
by case: PGFbFa => + _ _; apply; exact: decreasing_image_oo.
@@ -886,7 +887,8 @@ rewrite oppeD//= -(continuous_FTC2 ab _ _ DPGFE); last 2 first.
886887
- have [/= dF rF lF] := Fab.
887888
have := derivable_oo_continuous_bnd_within PGFbFa.
888889
move=> /(continuous_within_itvP _ FbFa)[_ PGFb PGFa]; split => /=.
889-
- move=> x xab; apply/derivable1_diffP; apply: differentiable_comp => //.
890+
- move=> x xab; apply/derivable1_diffP.
891+
apply: (differentiable_comp (g:[email protected] R)) => //.
890892
apply: differentiable_comp; apply/derivable1_diffP.
891893
by case: Fab => + _ _; exact.
892894
by case: PGFbFa => + _ _; apply; exact: decreasing_image_oo.
@@ -909,7 +911,7 @@ rewrite oppeD//= -(continuous_FTC2 ab _ _ DPGFE); last 2 first.
909911
rewrite gtr0_norm// ?subr_gt0//.
910912
by near: t; exact: nbhs_left_ltBl.
911913
apply: eq_integral_itv_bounded.
912-
- rewrite mulrN; apply: measurableT_comp => //.
914+
- rewrite mulrN; apply: (measurableT_comp (f:[email protected] R)) => //.
913915
apply: (eq_measurable_fun ((G \o F) * F^`())%R) => //.
914916
by move=> x; rewrite inE/= => xab; rewrite !fctE fE.
915917
by move: mGFF'; apply: measurable_funS => //; exact: subset_itv_oo_cc.
@@ -994,13 +996,14 @@ have mF' : measurable_fun `]a, b[ F^`().
994996
apply: subspace_continuous_measurable_fun => //.
995997
by apply: continuous_in_subspaceT => x /[!inE] xab; exact: cF'.
996998
rewrite integral_itv_bndoo//; last first.
997-
rewrite compA -(compA G -%R) (_ : -%R \o -%R = id); last first.
999+
rewrite (compA _ -%R) -(compA G -%R) (_ : -%R \o -%R = id); last first.
9981000
by apply/funext => y; rewrite /= opprK.
999-
apply: measurable_funM => //; apply: measurableT_comp => //.
1001+
apply: measurable_funM => //.
1002+
apply: (measurableT_comp (f:[email protected] R)) => //.
10001003
apply: (@eq_measurable_fun _ _ _ _ _ (- F^`())%R).
10011004
move=> x /[!inE] xab; rewrite [in RHS]derive1E deriveN -?derive1E//.
10021005
by case: Fab => + _ _; apply.
1003-
exact: measurableT_comp.
1006+
exact: (measurableT_comp (f:[email protected] R)).
10041007
rewrite [in RHS]integral_itv_bndoo//; last exact: measurable_funM.
10051008
apply: eq_integral => x /[!inE] xab; rewrite !fctE !opprK derive1E deriveN.
10061009
- by rewrite opprK -derive1E.

0 commit comments

Comments
 (0)