Minimum and maximum w.r.t. a filter and on a set #
Main Definitions #
This file defines six predicates of the form isAB, where A is Min, Max, or Extr,
and B is Filter or On.
- isMinFilter f l ameans that- f a ≤ f xin some- l-neighborhood of- a;
- isMaxFilter f l ameans that- f x ≤ f ain some- l-neighborhood of- a;
- isExtrFilter f l ameans- isMinFilter f l aor- isMaxFilter f l a.
Similar predicates with on suffix are particular cases for l = 𝓟 s.
Main statements #
Change of the filter (set) argument #
- is*Filter.filter_mono: replace the filter with a smaller one;
- is*Filter.filter_inf: replace a filter- lwith- l ⊓ l';
- is*On.on_subset: restrict to a smaller set;
- is*Pn.inter: replace a set- swith- s ∩ t.
Composition #
- is**.comp_mono: if- xis an extremum for- fand- gis a monotone function, then- xis an extremum for- g ∘ f;
- is**.comp_antitone: similarly for the case of antitone- g;
- is**.bicomp_mono: if- xis an extremum of the same type for- fand- gand a binary operation- opis monotone in both arguments, then- xis an extremum of the same type for- fun x => op (f x) (g x).
- is*Filter.comp_tendsto: if- g xis an extremum for- fw.r.t.- l'and- Tendsto g l l', then- xis an extremum for- f ∘ gw.r.t.- l.
- is*On.on_preimage: if- g xis an extremum for- fon- s, then- xis an extremum for- f ∘ gon- g ⁻¹' s.
Algebraic operations #
- is**.add: if- xis an extremum of the same type for two functions, then it is an extremum of the same type for their sum;
- is**.neg: if- xis an extremum for- f, then it is an extremum of the opposite type for- -f;
- is**.sub: if- xis a minimum for- fand a maximum for- g, then it is a minimum for- f - gand a maximum for- g - f;
- is**.max,- is**.min,- is**.sup,- is**.inf: similarly for- is**.addfor pointwise- max,- min,- sup,- inf, respectively.
Miscellaneous definitions #
- is**_const: any point is both a minimum and maximum for a constant function;
- isMin/Max*.isExt: any minimum/maximum point is an extremum;
- is**.dual,- is**.undual: conversion between codomains- αand- dual α;
Missing features (TODO) #
- Multiplication and division;
- is**.bicompl: if- xis a minimum for- f,- yis a minimum for- g, and- opis a monotone binary operation, then- (x, y)is a minimum for- uncurry (bicompl op f g). From this point of view,- is**.bicompis a composition
- It would be nice to have a tactic that specializes comp_(anti)monoorbicomp_monobased on a proof of monotonicity of a given (binary) function. The tactic should maintain ametalist of known (anti)monotone (binary) functions with their names, as well as a list of special types of filters, and define the missing lemmas once one of these two lists grows.
Definitions #
IsMinFilter f l a means that f a ≤ f x for all x in some l-neighborhood of a
Equations
- IsMinFilter f l a = ∀ᶠ (x : α) in l, f a ≤ f x
Instances For
is_maxFilter f l a means that f x ≤ f a for all x in some l-neighborhood of a
Equations
- IsMaxFilter f l a = ∀ᶠ (x : α) in l, f x ≤ f a
Instances For
IsExtrFilter f l a means IsMinFilter f l a or IsMaxFilter f l a
Equations
- IsExtrFilter f l a = (IsMinFilter f l a ∨ IsMaxFilter f l a)
Instances For
Conversion to IsExtr* #
Constant function #
Order dual #
Alias of the reverse direction of isMinFilter_dual_iff.
Alias of the forward direction of isMinFilter_dual_iff.
Alias of the forward direction of isMaxFilter_dual_iff.
Alias of the reverse direction of isMaxFilter_dual_iff.
Alias of the reverse direction of isExtrFilter_dual_iff.
Alias of the forward direction of isExtrFilter_dual_iff.
Alias of the forward direction of isMinOn_dual_iff.
Alias of the reverse direction of isMinOn_dual_iff.
Alias of the forward direction of isMaxOn_dual_iff.
Alias of the reverse direction of isMaxOn_dual_iff.
Alias of the forward direction of isExtrOn_dual_iff.
Alias of the reverse direction of isExtrOn_dual_iff.