mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-22 19:21:39 +03:00
[ funext ] Add a proof for funext variants with the other quantities
This commit is contained in:
parent
2a3f0311f2
commit
cf68e995c4
@ -127,6 +127,9 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
|
|||||||
|
|
||||||
* Removed need for the runtime value of the implicit argument in `succNotLTEpred`.
|
* Removed need for the runtime value of the implicit argument in `succNotLTEpred`.
|
||||||
|
|
||||||
|
* Added `funExt0` and `funExt1`, functions analogous to `funExt` but for functions
|
||||||
|
with quantities 0 and 1 respectively.
|
||||||
|
|
||||||
#### Contrib
|
#### Contrib
|
||||||
|
|
||||||
* `Data.List.Lazy` was moved from `contrib` to `base`.
|
* `Data.List.Lazy` was moved from `contrib` to `base`.
|
||||||
|
@ -10,3 +10,11 @@ module Control.Function.FunExt
|
|||||||
public export
|
public export
|
||||||
interface FunExt where
|
interface FunExt where
|
||||||
funExt : {0 b : a -> Type} -> {0 f, g : (x : a) -> b x} -> ((x : a) -> f x = g x) -> f = g
|
funExt : {0 b : a -> Type} -> {0 f, g : (x : a) -> b x} -> ((x : a) -> f x = g x) -> f = g
|
||||||
|
|
||||||
|
export
|
||||||
|
funExt0 : FunExt => {0 b : a -> Type} -> {0 f, g : (0 x : a) -> b x} -> ((x : a) -> f x = g x) -> f = g
|
||||||
|
funExt0 prf = rewrite funExt {f = \x => f x} {g = \y => g y} prf in Refl
|
||||||
|
|
||||||
|
export
|
||||||
|
funExt1 : FunExt => {0 b : a -> Type} -> {0 f, g : (1 x : a) -> b x} -> ((x : a) -> f x = g x) -> f = g
|
||||||
|
funExt1 prf = rewrite funExt {f = \x => f x} {g = \y => g y} prf in Refl
|
||||||
|
Loading…
Reference in New Issue
Block a user