mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-28 19:13:02 +03:00
Try a fix for the preorder reasoning error
Correct the dependency in `FastDerivation` to use the RHS in `step` Try to avoid shadowing `y` in the LHS of `Calc`
This commit is contained in:
parent
d387c76e9b
commit
ff76dee6c7
@ -19,12 +19,12 @@ public export
|
||||
public export
|
||||
data FastDerivation : (x : a) -> (y : b) -> Type where
|
||||
(|~) : (x : a) -> FastDerivation x x
|
||||
(~~) : FastDerivation x y -> (step : (z : c ** y ~=~ z)) -> FastDerivation x z
|
||||
(~~) : FastDerivation x y -> (step : (z : c ** y ~=~ z)) -> FastDerivation x (fst step)
|
||||
|
||||
public export
|
||||
Calc : {x : a} -> {y : b} -> FastDerivation x y -> x ~=~ y
|
||||
Calc (|~ x) = Refl
|
||||
Calc {y} ((~~) {z=y} {y=y} der (y ** Refl)) = Calc der
|
||||
Calc {y} ((~~) {z=_} {y=_} der (_ ** Refl)) = Calc der
|
||||
|
||||
{- -- requires import Data.Nat
|
||||
0
|
||||
|
Loading…
Reference in New Issue
Block a user