mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-24 06:52:19 +03:00
Erase additional PreorderReasoning arguments
This commit is contained in:
parent
dc43ce2056
commit
3c61471da3
@ -22,9 +22,12 @@ data FastDerivation : (x : a) -> (y : b) -> Type where
|
||||
(~~) : FastDerivation x y -> (step : Step y z) -> FastDerivation x z
|
||||
|
||||
public export
|
||||
Calc : {0 x : a} -> {0 y : b} -> FastDerivation x y -> x ~=~ y
|
||||
Calc (|~ x) = Refl
|
||||
Calc ((~~) der (_ ...(Refl))) = Calc der
|
||||
Calc : {0 x : a} -> {0 y : b} -> (0 der : FastDerivation x y) -> x ~=~ y
|
||||
Calc der = irrelevantEq $ Calc' der
|
||||
where
|
||||
Calc' : {0 x : c} -> {0 y : d} -> FastDerivation x y -> x ~=~ y
|
||||
Calc' (|~ x) = Refl
|
||||
Calc' ((~~) der (_ ...(Refl))) = Calc' der
|
||||
|
||||
{- -- requires import Data.Nat
|
||||
0
|
||||
@ -36,18 +39,10 @@ example x =
|
||||
~~ 1+x ...( plusCommutative x 1 )
|
||||
-}
|
||||
|
||||
-- Re-implement `prelude.Builtins.sym` to enable bootstrapping
|
||||
-- Remove after next release
|
||||
%inline
|
||||
public export
|
||||
symHet : (0 rule : x ~=~ y) -> y ~=~ x
|
||||
symHet Refl = Refl
|
||||
|
||||
-- Smart constructors
|
||||
public export
|
||||
(..<) : (0 y : a) -> {0 x : b} ->
|
||||
y ~=~ x -> Step x y
|
||||
(y ..<(prf)) {x} = (y ...(symHet prf)) -- Use `sym` from next release
|
||||
(..<) : (0 y : a) -> {0 x : b} -> (0 step : y ~=~ x) -> Step x y
|
||||
(y ..<(prf)) {x} = (y ...(sym prf))
|
||||
|
||||
public export
|
||||
(..>) : (0 y : a) -> (0 step : x ~=~ y) -> Step x y
|
||||
@ -55,5 +50,5 @@ public export
|
||||
|
||||
||| Use a judgemental equality but is not trivial to the reader.
|
||||
public export
|
||||
(~=) : FastDerivation x y -> (0 z : dom) -> {auto xEqy : y = z} -> FastDerivation x z
|
||||
(~=) : FastDerivation x y -> (0 z : dom) -> {auto 0 xEqy : y = z} -> FastDerivation x z
|
||||
(~=) der y {xEqy = Refl} = der ~~ y ... Refl
|
||||
|
Loading…
Reference in New Issue
Block a user