mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-28 02:23:44 +03:00
improvement: use std lib WellFounded
for Sufficient
This commit is contained in:
parent
18e887389f
commit
5fd5b1e732
@ -59,6 +59,20 @@ accInd : {0 rel : a -> a -> Type} -> {0 P : a -> Type} ->
|
|||||||
accInd step z (Access f) =
|
accInd step z (Access f) =
|
||||||
step z $ \y, lt => accInd step y (f y lt)
|
step z $ \y, lt => accInd step y (f y lt)
|
||||||
|
|
||||||
|
||| Depedently-typed induction for creating extrinsic proofs on results of `accInd`.
|
||||||
|
export
|
||||||
|
accIndProp : {0 P : a -> Type} ->
|
||||||
|
(step : (x : a) -> ((y : a) -> rel y x -> P y) -> P x) ->
|
||||||
|
{0 RP : (x : a) -> P x -> Type} ->
|
||||||
|
(ih : (x : a) ->
|
||||||
|
(f : (y : a) -> rel y x -> P y) ->
|
||||||
|
((y : a) -> (isRel : rel y x) -> RP y (f y isRel)) ->
|
||||||
|
RP x (step x f)) ->
|
||||||
|
(z : a) -> (0 acc : Accessible rel z) -> RP z (accInd step z acc)
|
||||||
|
accIndProp step ih z (Access rec) =
|
||||||
|
ih z (\y, lt => accInd step y (rec y lt))
|
||||||
|
(\y, lt => accIndProp {RP} step ih y (rec y lt))
|
||||||
|
|
||||||
||| Simply-typed recursion based on well-founded-ness.
|
||| Simply-typed recursion based on well-founded-ness.
|
||||||
|||
|
|||
|
||||||
||| This is `accRec` applied to accessibility derived from a `WellFounded`
|
||| This is `accRec` applied to accessibility derived from a `WellFounded`
|
||||||
@ -79,6 +93,19 @@ wfInd : (0 _ : WellFounded a rel) => {0 P : a -> Type} ->
|
|||||||
(myz : a) -> P myz
|
(myz : a) -> P myz
|
||||||
wfInd step myz = accInd step myz (wellFounded {rel} myz)
|
wfInd step myz = accInd step myz (wellFounded {rel} myz)
|
||||||
|
|
||||||
|
||| Depedently-typed induction for creating extrinsic proofs on results of `wfInd`.
|
||||||
|
export
|
||||||
|
wfIndProp : (0 _ : WellFounded a rel) =>
|
||||||
|
{0 P : a -> Type} ->
|
||||||
|
(step : (x : a) -> ((y : a) -> rel y x -> P y) -> P x) ->
|
||||||
|
{0 RP : (x : a) -> P x -> Type} ->
|
||||||
|
(ih : (x : a) ->
|
||||||
|
(f : (y : a) -> rel y x -> P y) ->
|
||||||
|
((y : a) -> (isRel : rel y x) -> RP y (f y isRel)) ->
|
||||||
|
RP x (step x f)) ->
|
||||||
|
(myz : a) -> RP myz (wfInd step myz)
|
||||||
|
wfIndProp step ih myz = accIndProp {RP} step ih myz (wellFounded {rel} myz)
|
||||||
|
|
||||||
||| Types that have a concept of size. The size must be a natural number.
|
||| Types that have a concept of size. The size must be a natural number.
|
||||||
public export
|
public export
|
||||||
interface Sized a where
|
interface Sized a where
|
||||||
|
@ -1,32 +1,23 @@
|
|||||||
||| 'Sufficient' lists: a structurally inductive view of lists xs
|
||| WellFounded on List suffixes
|
||||||
||| as given by xs = non-empty prefix + sufficient suffix
|
|
||||||
|||
|
|
||||||
||| Useful for termination arguments for function definitions
|
|
||||||
||| which provably consume a non-empty (but otherwise arbitrary) prefix
|
|
||||||
||| *without* having to resort to ancillary WF induction on length etc.
|
|
||||||
||| e.g. lexers, parsers etc.
|
|
||||||
|||
|
|
||||||
||| Credited by Conor McBride as originally due to James McKinna
|
|
||||||
module Data.List.Sufficient
|
module Data.List.Sufficient
|
||||||
|
|
||||||
||| Sufficient view
|
import Control.WellFounded
|
||||||
public export
|
|
||||||
data Sufficient : (xs : List a) -> Type where
|
|
||||||
SuffAcc : {xs : List a}
|
|
||||||
-> (suff_ih : {x : a} -> {pre, suff : List a}
|
|
||||||
-> xs = x :: (pre ++ suff)
|
|
||||||
-> Sufficient suff)
|
|
||||||
-> Sufficient xs
|
|
||||||
|
|
||||||
||| Sufficient view covering property
|
%default total
|
||||||
export
|
|
||||||
sufficient : (xs : List a) -> Sufficient xs
|
public export
|
||||||
sufficient [] = SuffAcc (\case _ impossible)
|
data Suffix : (ys,xs : List a) -> Type where
|
||||||
sufficient (x :: xs) with (sufficient xs)
|
IsSuffix : (x : a) -> (zs : List a)
|
||||||
sufficient (x :: xs) | suffxs@(SuffAcc suff_ih)
|
-> (0 ford : xs = x :: zs ++ ys) -> Suffix ys xs
|
||||||
= SuffAcc (\case Refl => prf Refl)
|
|
||||||
where prf : {pre, suff : List a}
|
SuffixAccessible : (xs : List a) -> Accessible Suffix xs
|
||||||
-> xs = pre ++ suff
|
SuffixAccessible [] = Access (\y => \case (IsSuffix x zs _) impossible)
|
||||||
-> Sufficient suff
|
SuffixAccessible ws@(x :: xs) =
|
||||||
prf {pre = []} Refl = suffxs
|
let fact1@(Access f) = SuffixAccessible xs
|
||||||
prf {pre = (y :: ys)} eq = suff_ih eq
|
in Access $ \ys => \case
|
||||||
|
(IsSuffix x [] Refl) => fact1
|
||||||
|
(IsSuffix x (z :: zs) Refl) => f ys (IsSuffix z zs Refl)
|
||||||
|
|
||||||
|
public export
|
||||||
|
WellFounded (List a) Suffix where
|
||||||
|
wellFounded = SuffixAccessible
|
||||||
|
Loading…
Reference in New Issue
Block a user