2020-10-06 15:09:02 +03:00
|
|
|
module Syntax.PreorderReasoning.Generic
|
|
|
|
|
2021-07-09 11:06:27 +03:00
|
|
|
import Control.Relation
|
|
|
|
import Control.Order
|
2023-02-17 20:47:54 +03:00
|
|
|
infixl 0 ~~, ~=
|
2020-10-06 15:09:02 +03:00
|
|
|
infixl 0 <~
|
|
|
|
prefix 1 |~
|
2022-03-22 23:58:36 +03:00
|
|
|
infix 1 ...,..<,..>,.=.,.=<,.=>
|
2020-10-06 15:09:02 +03:00
|
|
|
|
|
|
|
public export
|
|
|
|
data Step : (leq : a -> a -> Type) -> a -> a -> Type where
|
2020-10-16 00:44:30 +03:00
|
|
|
(...) : (y : a) -> x `leq` y -> Step leq x y
|
2020-10-06 15:09:02 +03:00
|
|
|
|
|
|
|
public export
|
2020-10-16 00:44:30 +03:00
|
|
|
data FastDerivation : (leq : a -> a -> Type) -> (x : a) -> (y : a) -> Type where
|
|
|
|
(|~) : (x : a) -> FastDerivation leq x x
|
|
|
|
(<~) : {x, y : a}
|
2023-02-17 20:47:54 +03:00
|
|
|
-> FastDerivation leq x y -> {0 z : a} -> (step : Step leq y z)
|
2020-10-16 00:44:30 +03:00
|
|
|
-> FastDerivation leq x z
|
2020-10-06 15:09:02 +03:00
|
|
|
|
2023-02-17 20:47:54 +03:00
|
|
|
-- Using a McKinna-McBride view records the depenencies on the type of derivation
|
|
|
|
-- the additional benefits for pattern-matching aren't used here
|
2020-10-16 00:44:30 +03:00
|
|
|
public export
|
2023-02-17 20:47:54 +03:00
|
|
|
data DerivationType : FastDerivation leq x y -> Type where
|
|
|
|
TrivialDerivation : DerivationType (|~ x)
|
|
|
|
SingleStepDerivation : DerivationType (|~ x <~ step)
|
|
|
|
NonTrivialDerivation : DerivationType (der <~ step1 <~ step2)
|
|
|
|
|
|
|
|
public export
|
|
|
|
derivationType : (der : FastDerivation leq x y) -> DerivationType der
|
|
|
|
derivationType (|~ x) = TrivialDerivation
|
|
|
|
derivationType (|~ x <~ step) = SingleStepDerivation
|
|
|
|
derivationType ((z <~ w) <~ step) = NonTrivialDerivation
|
|
|
|
|
|
|
|
public export
|
|
|
|
0 Prerequisite : {0 der : FastDerivation {a = dom} leq x y} -> DerivationType der -> Type
|
|
|
|
Prerequisite TrivialDerivation = Reflexive dom leq
|
|
|
|
Prerequisite SingleStepDerivation = ()
|
|
|
|
Prerequisite NonTrivialDerivation = Transitive dom leq
|
|
|
|
|
|
|
|
-- Once we have the prerequisite for transitivity, we have the
|
|
|
|
-- prerequisite for its inductive predecessor:
|
|
|
|
public export
|
|
|
|
inductivePrerequisite : (der : FastDerivation leq x y) ->
|
|
|
|
(0 step1 : Step leq y z) -> (0 step2 : Step leq z w) ->
|
|
|
|
Prerequisite (derivationType (der <~ step1 <~ step2)) ->
|
|
|
|
Prerequisite (derivationType (der <~ step1))
|
|
|
|
inductivePrerequisite ( |~ y ) step1 step2 prerequisite = ()
|
|
|
|
inductivePrerequisite (z <~ step0) step1 step2 prerequisite = prerequisite
|
|
|
|
|
|
|
|
public export
|
|
|
|
preorderPrerequisite : Preorder dom leq => (der : FastDerivation leq x y) -> Prerequisite (derivationType der)
|
|
|
|
preorderPrerequisite (|~ x) = %search
|
|
|
|
preorderPrerequisite (|~ x <~ step) = ()
|
|
|
|
preorderPrerequisite (der <~ step1 <~ step2) = %search
|
|
|
|
|
|
|
|
||| The Prerequisite for the derivation:
|
|
|
|
||| 0-length derivation: Reflexive dom leq
|
|
|
|
||| 1-length derivation: Unit (no prerequisite)
|
|
|
|
||| 2 steps of longer: Transitivity
|
|
|
|
public export
|
|
|
|
CalcSmart : {0 x : dom} -> {0 y : dom} -> (der : FastDerivation leq x y) ->
|
|
|
|
Prerequisite (derivationType der) => x `leq` y
|
|
|
|
CalcSmart (|~ x) = reflexive
|
|
|
|
CalcSmart (|~ x <~ y ... step) = step
|
|
|
|
CalcSmart (der <~ _ ... step1
|
|
|
|
<~ w ... step2) @{prereq} =
|
|
|
|
CalcSmart (der <~ _ ... step1)
|
|
|
|
@{inductivePrerequisite der (_ ... step1) (w ... step2) prereq}
|
|
|
|
`transitive` step2
|
|
|
|
|
|
|
|
public export
|
|
|
|
CalcWith : Preorder dom leq => {0 x : dom} -> {0 y : dom} -> (der : FastDerivation leq x y) -> x `leq` y
|
|
|
|
CalcWith der = CalcSmart der @{preorderPrerequisite der}
|
2020-10-06 15:09:02 +03:00
|
|
|
|
|
|
|
public export
|
2021-02-12 18:36:38 +03:00
|
|
|
(~~) : {0 x : dom} -> {0 y : dom}
|
2023-02-17 20:47:54 +03:00
|
|
|
-> FastDerivation leq x y -> {0 z : dom} -> (0 step : Step Equal y z)
|
2020-10-16 00:44:30 +03:00
|
|
|
-> FastDerivation leq x z
|
2023-02-17 20:47:54 +03:00
|
|
|
(~~) der (z ... yEqZ) = replace {p = FastDerivation leq _} yEqZ der
|
2022-03-22 23:58:36 +03:00
|
|
|
|
|
|
|
-- Smart constructors
|
|
|
|
public export
|
|
|
|
(..<) : Symmetric a leq => (y : a) -> {x : a} -> y `leq` x -> Step leq x y
|
|
|
|
(y ..<(prf)) {x} = (y ...(symmetric prf))
|
|
|
|
|
|
|
|
public export
|
|
|
|
(..>) : (y : a) -> x `leq` y -> Step leq x y
|
|
|
|
(..>) = (...)
|
|
|
|
|
|
|
|
public export
|
|
|
|
(.=.) : Reflexive a leq => (y : a) -> {x : a} ->
|
|
|
|
x === y -> Step leq x y
|
|
|
|
(y .=.(Refl)) {x = _} = (y ...(reflexive))
|
|
|
|
|
|
|
|
public export
|
|
|
|
(.=>) : Reflexive a leq => (y : a) -> {x : a} ->
|
|
|
|
x === y -> Step leq x y
|
|
|
|
(.=>) = (.=.)
|
|
|
|
|
|
|
|
public export
|
|
|
|
(.=<) : Reflexive a leq => (y : a) -> {x : a} ->
|
|
|
|
y === x -> Step leq x y
|
|
|
|
(y .=<(Refl)) = (y ...(reflexive))
|
2023-02-17 20:47:54 +03:00
|
|
|
|
|
|
|
public export
|
|
|
|
(~=) : FastDerivation leq x y -> (0 z : dom) -> {auto xEqy : y = z} -> FastDerivation leq x z
|
|
|
|
(~=) der z {xEqy} = der ~~ z ...(xEqy)
|