fun_prop
environment extensions storing theorems for fun_prop
#
Stores important argument indices of lambda theorems
For example
theorem Continuous_const {α β} [TopologicalSpace α] [TopologicalSpace β] (y : β) :
Continuous fun _ : α => y
is represented by
.const 0 4
- id: ℕ → Mathlib.Meta.FunProp.LambdaTheoremArgs
- const: ℕ → ℕ → Mathlib.Meta.FunProp.LambdaTheoremArgs
- proj: ℕ → ℕ → Mathlib.Meta.FunProp.LambdaTheoremArgs
- projDep: ℕ → ℕ → Mathlib.Meta.FunProp.LambdaTheoremArgs
- comp: ℕ → ℕ → Mathlib.Meta.FunProp.LambdaTheoremArgs
- letE: ℕ → ℕ → Mathlib.Meta.FunProp.LambdaTheoremArgs
- pi: ℕ → Mathlib.Meta.FunProp.LambdaTheoremArgs
Instances For
There are 5(+1) basic lambda theorems
- id
Continuous fun x => x
- const
Continuous fun x => y
- proj
Continuous fun (f : X → Y) => f x
- projDep
Continuous fun (f : (x : X) → Y x => f x)
- comp
Continuous f → Continuous g → Continuous fun x => f (g x)
- letE
Continuous f → Continuous g → Continuous fun x => let y := g x; f x y
- pi
∀ y, Continuous (f · y) → Continuous fun x y => f x y
- id: Mathlib.Meta.FunProp.LambdaTheoremType
- const: Mathlib.Meta.FunProp.LambdaTheoremType
- proj: Mathlib.Meta.FunProp.LambdaTheoremType
- projDep: Mathlib.Meta.FunProp.LambdaTheoremType
- comp: Mathlib.Meta.FunProp.LambdaTheoremType
- letE: Mathlib.Meta.FunProp.LambdaTheoremType
- pi: Mathlib.Meta.FunProp.LambdaTheoremType
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
- funPropName : Lean.Name
Name of function property
- thmName : Lean.Name
Name of lambda theorem
- thmArgs : Mathlib.Meta.FunProp.LambdaTheoremArgs
Type and important argument of the theorem.
Instances For
Equations
- Mathlib.Meta.FunProp.instInhabitedLambdaTheorem = { default := { funPropName := default, thmName := default, thmArgs := default } }
- theorems : Lean.HashMap (Lean.Name × Mathlib.Meta.FunProp.LambdaTheoremType) Mathlib.Meta.FunProp.LambdaTheorem
map: function property name × theorem type → lambda theorem
Instances For
Equations
- Mathlib.Meta.FunProp.instInhabitedLambdaTheorems = { default := { theorems := default } }
return proof of lambda theorem
Equations
Instances For
Equations
Instances For
Extension storing all lambda theorems.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Function theorems are stated in uncurried or compositional form.
uncurried
theorem Continuous_add : Continuous (fun x => x.1 + x.2)
compositional
theorem Continuous_add (hf : Continuous f) (hg : Continuous g) : Continuous (fun x => (f x) + (g x))
- uncurried: Mathlib.Meta.FunProp.TheoremForm
- comp: Mathlib.Meta.FunProp.TheoremForm
Instances For
theorem about specific function (either declared constant or free variable)
- funPropName : Lean.Name
function property name
- thmOrigin : Mathlib.Meta.FunProp.Origin
theorem name
- funOrigin : Mathlib.Meta.FunProp.Origin
function name
array of argument indices about which this theorem is about
- appliedArgs : ℕ
total number of arguments applied to the function
- priority : ℕ
priority
form of the theorem, see documentation of TheoremForm
Instances For
Equations
- One or more equations did not get rendered due to their size.
Equations
- Mathlib.Meta.FunProp.instInhabitedFunctionTheorems = { default := { theorems := default } }
return proof of function theorem
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
Instances For
Extension storing all function theorems.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Mathlib.Meta.FunProp.instInhabitedGeneralTheorem = { default := { funPropName := default, thmName := default, keys := default, priority := default } }
Get proof of a theorem.
Equations
Instances For
Instances For
Equations
- Mathlib.Meta.FunProp.instInhabitedGeneralTheorems = { default := { theorems := default } }
Equations
Instances For
There are four types of theorems:
- lam - theorem about basic lambda calculus terms
- function - theorem about a specific function(declared or free variable) in specific arguments
- mor - special theorems talking about bundled morphisms/DFunLike.coe
- transition - theorems inferring one function property from another
Examples:
- lam
theorem Continuous_id : Continuous fun x => x
theorem Continuous_comp (hf : Continuous f) (hg : Continuous g) : Continuous fun x => f (g x)
- function
theorem Continuous_add : Continuous (fun x => x.1 + x.2)
theorem Continuous_add (hf : Continuous f) (hg : Continuous g) :
Continuous (fun x => (f x) + (g x))
- mor - the head of function body has to be ``DFunLike.code
theorem ContDiff.clm_apply {f : E → F →L[𝕜] G} {g : E → F}
(hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) :
ContDiff 𝕜 n fun x => (f x) (g x)
theorem clm_linear {f : E →L[𝕜] F} : IsLinearMap 𝕜 f
- transition - the conclusion has to be in the form
P f
wheref
is a free variable
theorem linear_is_continuous [FiniteDimensional ℝ E] {f : E → F} (hf : IsLinearMap 𝕜 f) :
Continuous f
- lam: Mathlib.Meta.FunProp.LambdaTheorem → Mathlib.Meta.FunProp.Theorem
- function: Mathlib.Meta.FunProp.FunctionTheorem → Mathlib.Meta.FunProp.Theorem
- mor: Mathlib.Meta.FunProp.GeneralTheorem → Mathlib.Meta.FunProp.Theorem
- transition: Mathlib.Meta.FunProp.GeneralTheorem → Mathlib.Meta.FunProp.Theorem
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.