mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 12:42:02 +03:00
Add symmetric and transitive closure relations
This commit is contained in:
parent
3c9393e5a8
commit
9bfa04921a
@ -127,3 +127,59 @@ PartialEquivalence ty Equal where
|
|||||||
|
|
||||||
public export
|
public export
|
||||||
Equivalence ty Equal where
|
Equivalence ty Equal where
|
||||||
|
|
||||||
|
----------------------------------------
|
||||||
|
|
||||||
|
public export
|
||||||
|
data SymClosure : Rel ty -> ty -> ty -> Type where
|
||||||
|
Fwd : {0 x, y : ty} -> rel x y -> SymClosure rel x y
|
||||||
|
Bwd : {0 x, y : ty} -> rel y x -> SymClosure rel x y
|
||||||
|
|
||||||
|
public export
|
||||||
|
Reflexive ty rel => Reflexive ty (SymClosure rel) where
|
||||||
|
reflexive = Fwd reflexive
|
||||||
|
|
||||||
|
public export
|
||||||
|
Symmetric ty (SymClosure rel) where
|
||||||
|
symmetric (Fwd xy) = Bwd xy
|
||||||
|
symmetric (Bwd yx) = Fwd yx
|
||||||
|
|
||||||
|
----------------------------------------
|
||||||
|
|
||||||
|
public export
|
||||||
|
data TransClosure : Rel ty -> ty -> ty -> Type where
|
||||||
|
Nil : TransClosure rel x x
|
||||||
|
(::) : {y : ty} -> rel x y -> TransClosure rel y z -> TransClosure rel x z
|
||||||
|
|
||||||
|
public export
|
||||||
|
Reflexive ty (TransClosure rel) where
|
||||||
|
reflexive = []
|
||||||
|
|
||||||
|
public export
|
||||||
|
Symmetric ty rel => Symmetric ty (TransClosure rel) where
|
||||||
|
symmetric {x} {y} xy = reverseOnto [] xy
|
||||||
|
where
|
||||||
|
reverseOnto : {z : ty} ->
|
||||||
|
TransClosure rel z x ->
|
||||||
|
TransClosure rel z y ->
|
||||||
|
TransClosure rel y x
|
||||||
|
reverseOnto zx [] = zx
|
||||||
|
reverseOnto zx (zw :: wy) = reverseOnto (symmetric zw :: zx) wy
|
||||||
|
|
||||||
|
public export
|
||||||
|
Transitive ty (TransClosure rel) where
|
||||||
|
transitive [] yz = yz
|
||||||
|
transitive (xw :: wy) yz = xw :: transitive wy yz
|
||||||
|
|
||||||
|
public export
|
||||||
|
Symmetric ty rel => Euclidean ty (TransClosure rel) where
|
||||||
|
euclidean = euclidean @{TSE}
|
||||||
|
|
||||||
|
public export
|
||||||
|
Symmetric ty rel => Tolerance ty (TransClosure rel) where
|
||||||
|
|
||||||
|
public export
|
||||||
|
Symmetric ty rel => PartialEquivalence ty (TransClosure rel) where
|
||||||
|
|
||||||
|
public export
|
||||||
|
Symmetric ty rel => Equivalence ty (TransClosure rel) where
|
||||||
|
@ -906,12 +906,12 @@ groupAllWith f = groupWith f . sortBy (comparing f)
|
|||||||
|
|
||||||
-- Nil is not Cons
|
-- Nil is not Cons
|
||||||
export
|
export
|
||||||
Uninhabited ([] = x :: xs) where
|
{0 xs : List a} -> Uninhabited ([] = x :: xs) where
|
||||||
uninhabited Refl impossible
|
uninhabited Refl impossible
|
||||||
|
|
||||||
-- Cons is not Nil
|
-- Cons is not Nil
|
||||||
export
|
export
|
||||||
Uninhabited (x :: xs = []) where
|
{0 xs : List a} -> Uninhabited (x :: xs = []) where
|
||||||
uninhabited Refl impossible
|
uninhabited Refl impossible
|
||||||
|
|
||||||
-- If the heads or the tails of two lists are provably non-equal, then the
|
-- If the heads or the tails of two lists are provably non-equal, then the
|
||||||
|
@ -109,21 +109,25 @@ LOG elab.ambiguous:5: Ambiguous elaboration at Vec:L:C--L:C:
|
|||||||
($resolvedN 2)
|
($resolvedN 2)
|
||||||
($resolvedN 2)
|
($resolvedN 2)
|
||||||
With default. Target type : Prelude.Types.Nat
|
With default. Target type : Prelude.Types.Nat
|
||||||
LOG elab.ambiguous:5: Ambiguous elaboration (kept 3 out of 3 candidates) (not delayed) at Vec:L:C--L:C:
|
LOG elab.ambiguous:5: Ambiguous elaboration (kept 4 out of 4 candidates) (not delayed) at Vec:L:C--L:C:
|
||||||
|
(($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil))
|
||||||
(($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil))
|
(($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil))
|
||||||
(($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil))
|
(($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil))
|
||||||
(($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil))
|
(($resolvedN Nil) ((:: ((:: (fromInteger 0)) Nil)) Nil))
|
||||||
Target type : ({arg:N} : (Data.Fin.Fin (Prelude.Types.S (Prelude.Types.S Prelude.Types.Z)))) -> (Prelude.Basics.List Prelude.Types.Nat)
|
Target type : ({arg:N} : (Data.Fin.Fin (Prelude.Types.S (Prelude.Types.S Prelude.Types.Z)))) -> (Prelude.Basics.List Prelude.Types.Nat)
|
||||||
LOG elab.ambiguous:5: Ambiguous elaboration (kept 2 out of 2 candidates) (not delayed) at Vec:L:C--L:C:
|
LOG elab.ambiguous:5: Ambiguous elaboration (kept 3 out of 3 candidates) (not delayed) at Vec:L:C--L:C:
|
||||||
|
$resolvedN
|
||||||
$resolvedN
|
$resolvedN
|
||||||
$resolvedN
|
$resolvedN
|
||||||
Target type : ?Vec.{a:N}_[]
|
Target type : ?Vec.{a:N}_[]
|
||||||
LOG elab.ambiguous:5: Ambiguous elaboration (kept 3 out of 3 candidates) (not delayed) at Vec:L:C--L:C:
|
LOG elab.ambiguous:5: Ambiguous elaboration (kept 4 out of 4 candidates) (not delayed) at Vec:L:C--L:C:
|
||||||
|
(($resolvedN ((:: (fromInteger 0)) Nil)) Nil)
|
||||||
(($resolvedN ((:: (fromInteger 0)) Nil)) Nil)
|
(($resolvedN ((:: (fromInteger 0)) Nil)) Nil)
|
||||||
(($resolvedN ((:: (fromInteger 0)) Nil)) Nil)
|
(($resolvedN ((:: (fromInteger 0)) Nil)) Nil)
|
||||||
(($resolvedN ((:: (fromInteger 0)) Nil)) Nil)
|
(($resolvedN ((:: (fromInteger 0)) Nil)) Nil)
|
||||||
Target type : (Vec.Vec ?Vec.{a:N}_[] ?Vec.{n:N}_[])
|
Target type : (Vec.Vec ?Vec.{a:N}_[] ?Vec.{n:N}_[])
|
||||||
LOG elab.ambiguous:5: Ambiguous elaboration (kept 3 out of 3 candidates) (not delayed) at Vec:L:C--L:C:
|
LOG elab.ambiguous:5: Ambiguous elaboration (kept 4 out of 4 candidates) (not delayed) at Vec:L:C--L:C:
|
||||||
|
(($resolvedN (fromInteger 0)) Nil)
|
||||||
(($resolvedN (fromInteger 0)) Nil)
|
(($resolvedN (fromInteger 0)) Nil)
|
||||||
(($resolvedN (fromInteger 0)) Nil)
|
(($resolvedN (fromInteger 0)) Nil)
|
||||||
(($resolvedN (fromInteger 0)) Nil)
|
(($resolvedN (fromInteger 0)) Nil)
|
||||||
@ -132,26 +136,32 @@ LOG elab.ambiguous:5: Ambiguous elaboration at Vec:L:C--L:C:
|
|||||||
($resolvedN 0)
|
($resolvedN 0)
|
||||||
($resolvedN 0)
|
($resolvedN 0)
|
||||||
With default. Target type : ?Vec.{a:N}_[]
|
With default. Target type : ?Vec.{a:N}_[]
|
||||||
LOG elab.ambiguous:5: Ambiguous elaboration (kept 2 out of 2 candidates) (not delayed) at Vec:L:C--L:C:
|
LOG elab.ambiguous:5: Ambiguous elaboration (kept 3 out of 3 candidates) (not delayed) at Vec:L:C--L:C:
|
||||||
|
$resolvedN
|
||||||
$resolvedN
|
$resolvedN
|
||||||
$resolvedN
|
$resolvedN
|
||||||
Target type : (Vec.Vec ?Vec.{a:N}_[] ?Vec.{n:N}_[])
|
Target type : (Vec.Vec ?Vec.{a:N}_[] ?Vec.{n:N}_[])
|
||||||
|
LOG elab.ambiguous:5: Ambiguous elaboration at Vec:L:C--L:C:
|
||||||
|
($resolvedN 0)
|
||||||
|
($resolvedN 0)
|
||||||
|
With default. Target type : (?Vec.{rel:N}_[] ?Vec.{x:N}_[] ?Vec.{y:N}_[])
|
||||||
LOG elab.ambiguous:5: Ambiguous elaboration at Vec:L:C--L:C:
|
LOG elab.ambiguous:5: Ambiguous elaboration at Vec:L:C--L:C:
|
||||||
($resolvedN 0)
|
($resolvedN 0)
|
||||||
($resolvedN 0)
|
($resolvedN 0)
|
||||||
With default. Target type : ?Vec.{a:N}_[]
|
With default. Target type : ?Vec.{a:N}_[]
|
||||||
LOG elab.ambiguous:5: Ambiguous elaboration (kept 2 out of 2 candidates) (not delayed) at Vec:L:C--L:C:
|
LOG elab.ambiguous:5: Ambiguous elaboration (kept 3 out of 3 candidates) (not delayed) at Vec:L:C--L:C:
|
||||||
|
$resolvedN
|
||||||
$resolvedN
|
$resolvedN
|
||||||
$resolvedN
|
$resolvedN
|
||||||
Target type : (Vec.Vec ?Vec.{a:N}_[] ?Vec.{n:N}_[])
|
Target type : (Vec.Vec ?Vec.{a:N}_[] ?Vec.{n:N}_[])
|
||||||
LOG elab.ambiguous:5: Ambiguous elaboration (kept 1 out of 3 candidates) (delayed) at Vec:L:C--L:C:
|
LOG elab.ambiguous:5: Ambiguous elaboration (kept 1 out of 4 candidates) (delayed) at Vec:L:C--L:C:
|
||||||
(($resolvedN (fromInteger 0)) Nil)
|
(($resolvedN (fromInteger 0)) Nil)
|
||||||
Target type : (Prelude.Basics.List Prelude.Types.Nat)
|
Target type : (Prelude.Basics.List Prelude.Types.Nat)
|
||||||
LOG elab.ambiguous:5: Ambiguous elaboration at Vec:L:C--L:C:
|
LOG elab.ambiguous:5: Ambiguous elaboration at Vec:L:C--L:C:
|
||||||
($resolvedN 0)
|
($resolvedN 0)
|
||||||
($resolvedN 0)
|
($resolvedN 0)
|
||||||
With default. Target type : Prelude.Types.Nat
|
With default. Target type : Prelude.Types.Nat
|
||||||
LOG elab.ambiguous:5: Ambiguous elaboration (kept 1 out of 2 candidates) (delayed) at Vec:L:C--L:C:
|
LOG elab.ambiguous:5: Ambiguous elaboration (kept 1 out of 3 candidates) (delayed) at Vec:L:C--L:C:
|
||||||
$resolvedN
|
$resolvedN
|
||||||
Target type : (Prelude.Basics.List Prelude.Types.Nat)
|
Target type : (Prelude.Basics.List Prelude.Types.Nat)
|
||||||
LOG declare.def:2: Case tree for Vec.test: [0] (Vec.:: (Prelude.Types.S Prelude.Types.Z) (Prelude.Basics.List Prelude.Types.Nat) (Prelude.Basics.Nil Prelude.Types.Nat) (Vec.:: Prelude.Types.Z (Prelude.Basics.List Prelude.Types.Nat) (Prelude.Basics.:: Prelude.Types.Nat Prelude.Types.Z (Prelude.Basics.Nil Prelude.Types.Nat)) (Vec.Nil (Prelude.Basics.List Prelude.Types.Nat))))
|
LOG declare.def:2: Case tree for Vec.test: [0] (Vec.:: (Prelude.Types.S Prelude.Types.Z) (Prelude.Basics.List Prelude.Types.Nat) (Prelude.Basics.Nil Prelude.Types.Nat) (Vec.:: Prelude.Types.Z (Prelude.Basics.List Prelude.Types.Nat) (Prelude.Basics.:: Prelude.Types.Nat Prelude.Types.Z (Prelude.Basics.Nil Prelude.Types.Nat)) (Vec.Nil (Prelude.Basics.List Prelude.Types.Nat))))
|
||||||
|
@ -14,6 +14,7 @@ Main> Error: Can't find an implementation for Num ().
|
|||||||
^^^^
|
^^^^
|
||||||
|
|
||||||
Main> Error: Ambiguous elaboration. Possible results:
|
Main> Error: Ambiguous elaboration. Possible results:
|
||||||
|
Control.Relation.Nil
|
||||||
Prelude.Nil
|
Prelude.Nil
|
||||||
Data.Vect.Nil
|
Data.Vect.Nil
|
||||||
|
|
||||||
|
Loading…
Reference in New Issue
Block a user