Skip to content

Make the property argument to replace explicit #3531

New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

Open
wants to merge 3 commits into
base: main
Choose a base branch
from
Open
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
2 changes: 2 additions & 0 deletions CHANGELOG_NEXT.md
Original file line number Diff line number Diff line change
Expand Up @@ -206,6 +206,8 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO

#### Prelude

* Made the property argument to `replace` explicit.

* Added pipeline operators `(|>)` and `(<|)`.

#### Base
Expand Down
10 changes: 5 additions & 5 deletions docs/source/proofs/patterns.rst
Original file line number Diff line number Diff line change
Expand Up @@ -180,16 +180,16 @@ prelude:
.. code-block:: idris

Main> :t replace
Builtin.replace : (0 rule : x = y) -> p x -> p y
Builtin.replace : (0 p : _) -> (0 rule : x = y) -> p x -> p y

Given a proof that ``x = y``, and a property ``p`` which holds for
``x``, we can get a proof of the same property for ``y``, because we
know ``x`` and ``y`` must be the same. Note the multiplicity on ``rule``
means that it's guaranteed to be erased at run time.
In practice, this function can be
a little tricky to use because in general the implicit argument ``p``
can be hard to infer by unification, so Idris provides a high level
syntax which calculates the property and applies ``replace``:
In practice, this function can be a little tricky to use because in
general the argument ``p`` can be hard to infer by unification, so
Idris provides a high level syntax which calculates the property
and applies ``replace``:
Comment on lines +189 to +192
Copy link
Contributor Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I don't quite follow the logic here, but I've kept it as similar as possible in wording. Do edit it if you know a better wording (also in the other file which duplicates this)

Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I don't think there is any "infer by unification" here anymore, since p was implicit and this PR makes it explicit. Maybe something along the lines of "the argument p is cumbersome to write, so Idris provides ....".

In practice, I use replace when Idris can't figure out p in rewrite, especially in the case where I only want to replace a subset of the occurrences of a term. But others who have more experience with using Idris for proofs than I do. I've always had to specify p when I use replace.


.. code-block:: idris

Expand Down
14 changes: 7 additions & 7 deletions docs/source/proofs/propositional.rst
Original file line number Diff line number Diff line change
Expand Up @@ -56,8 +56,8 @@ To use this in our proofs there is the following function in the prelude:
.. code-block:: idris

||| Perform substitution in a term according to some equality.
replace : forall x, y, prop . (0 rule : x = y) -> prop x -> prop y
replace Refl prf = prf
replace : forall x, y . (0 prop : _) -> (0 rule : x = y) -> prop x -> prop y
replace _ Refl prf = prf

If we supply an equality (x=y) and a proof of a property of x (``prop x``) then we get
a proof of a property of y (``prop y``).
Expand All @@ -70,15 +70,15 @@ the equality ``x=y`` then we get a proof that ``y=2``.
p1 n = (n=2)

testReplace: (x=y) -> (p1 x) -> (p1 y)
testReplace a b = replace a b
testReplace a b = replace p1 a b

Rewrite
=======

In practice, ``replace`` can be
a little tricky to use because in general the implicit argument ``prop``
can be hard to infer for the machine, so Idris provides a high level
syntax which calculates the property and applies ``replace``.
In practice, ``replace`` can be a little tricky to use because in
general the argument ``prop`` can be hard to infer for the machine,
so Idris provides a high level syntax which calculates the property
and applies ``replace``.

Example: again we supply ``p1 x`` which is a proof that ``x=2`` and the equality
``y=x`` then we get a proof that ``y=2``.
Expand Down
2 changes: 1 addition & 1 deletion docs/source/tutorial/theorems.rst
Original file line number Diff line number Diff line change
Expand Up @@ -60,7 +60,7 @@ never equal to a successor:
.. code-block:: idris

disjoint : (n : Nat) -> Z = S n -> Void
disjoint n prf = replace {p = disjointTy} prf ()
disjoint n prf = replace disjointTy prf ()
where
disjointTy : Nat -> Type
disjointTy Z = ()
Expand Down
2 changes: 1 addition & 1 deletion libs/base/Data/Fin.idr
Original file line number Diff line number Diff line change
Expand Up @@ -28,7 +28,7 @@ coerce {n = Z} eq FZ impossible
coerce {n = S _} eq (FS k) = FS (coerce (injective eq) k)
coerce {n = Z} eq (FS k) impossible

%transform "coerce-identity" coerce eq k = replace {p = Fin} eq k
%transform "coerce-identity" coerce eq k = replace Fin eq k

export
Uninhabited (Fin Z) where
Expand Down
6 changes: 3 additions & 3 deletions libs/base/Data/Nat/Order/Properties.idr
Original file line number Diff line number Diff line change
Expand Up @@ -26,15 +26,15 @@ ltReflection a = lteReflection (S a)
-- For example:
export
lteIsLTE : (a, b : Nat) -> a `lte` b = True -> a `LTE` b
lteIsLTE a b prf = invert (replace {p = Reflects (a `LTE` b)} prf (lteReflection a b))
lteIsLTE a b prf = invert (replace (Reflects (a `LTE` b)) prf (lteReflection a b))

export
ltIsLT : (a, b : Nat) -> a `lt` b = True -> a `LT` b
ltIsLT a = lteIsLTE (S a)

export
notlteIsNotLTE : (a, b : Nat) -> a `lte` b = False -> Not (a `LTE` b)
notlteIsNotLTE a b prf = invert (replace {p = Reflects (a `LTE` b)} prf (lteReflection a b))
notlteIsNotLTE a b prf = invert (replace (Reflects (a `LTE` b)) prf (lteReflection a b))

export
notltIsNotLT : (a, b : Nat) -> a `lt` b = False -> Not (a `LT` b)
Expand All @@ -45,7 +45,7 @@ export
notlteIsLT : (a, b : Nat) -> a `lte` b = False -> b `LT` a
notlteIsLT a b prf = notLTImpliesGTE $
\prf' =>
(invert $ replace {p = Reflects (S a `LTE` S b)} prf
(invert $ replace (Reflects (S a `LTE` S b)) prf
$ lteReflection (S a) (S b)) prf'

export
Expand Down
2 changes: 1 addition & 1 deletion libs/base/Syntax/PreorderReasoning/Generic.idr
Original file line number Diff line number Diff line change
Expand Up @@ -74,7 +74,7 @@ public export
(~~) : {0 x : dom} -> {0 y : dom}
-> FastDerivation leq x y -> {0 z : dom} -> (0 step : Step Equal y z)
-> FastDerivation leq x z
(~~) der (z ... yEqZ) = replace {p = FastDerivation leq _} yEqZ der
(~~) der (z ... yEqZ) = replace (FastDerivation leq _) yEqZ der

-- Smart constructors
public export
Expand Down
10 changes: 5 additions & 5 deletions libs/contrib/Data/Nat/Factor.idr
Original file line number Diff line number Diff line change
Expand Up @@ -199,7 +199,7 @@ export
factorLteNumber : Factor p n -> {auto positN : LTE 1 n} -> LTE p n
factorLteNumber (CofactorExists Z prf) =
let nIsZero = trans prf $ multCommutative p 0
oneLteZero = replace {p = LTE 1} nIsZero positN
oneLteZero = replace (LTE 1) nIsZero positN
in
absurd $ succNotLTEzero oneLteZero
factorLteNumber (CofactorExists (S k) prf) =
Expand Down Expand Up @@ -258,7 +258,7 @@ minusFactor (CofactorExists qab prfAB) (CofactorExists qa prfA) =
rewrite multDistributesOverMinusRight p qab qa in
rewrite sym prfA in
rewrite sym prfAB in
replace {p = \x => b = minus (a + b) x} (plusZeroRightNeutral a) $
replace (\x => b = minus (a + b) x) (plusZeroRightNeutral a) $
rewrite plusMinusLeftCancel a b 0 in
rewrite minusZeroRight b in
Refl
Expand Down Expand Up @@ -459,15 +459,15 @@ export
divByGcdGcdOne : {a, b, c : Nat} -> GCD a (a * b) (a * c) -> GCD 1 b c
divByGcdGcdOne {a = Z} (MkGCD _ _) impossible
divByGcdGcdOne {a = S a} {b = Z} {c = Z} (MkGCD {notBothZero} _ _) =
case replace {p = \x => NotBothZero x x} (multZeroRightZero (S a)) notBothZero of
case replace (\x => NotBothZero x x) (multZeroRightZero (S a)) notBothZero of
LeftIsNotZero impossible
RightIsNotZero impossible
divByGcdGcdOne {a = S a} {b = Z} {c = S c} gcdPrf@(MkGCD {notBothZero} _ _) =
case replace {p = \x => NotBothZero x (S a * S c)} (multZeroRightZero (S a)) notBothZero of
case replace (\x => NotBothZero x (S a * S c)) (multZeroRightZero (S a)) notBothZero of
LeftIsNotZero impossible
RightIsNotZero => symmetric $ divByGcdHelper a c Z $ symmetric gcdPrf
divByGcdGcdOne {a = S a} {b = S b} {c = Z} gcdPrf@(MkGCD {notBothZero} _ _) =
case replace {p = \x => NotBothZero (S a * S b) x} (multZeroRightZero (S a)) notBothZero of
case replace (\x => NotBothZero (S a * S b) x) (multZeroRightZero (S a)) notBothZero of
RightIsNotZero impossible
LeftIsNotZero => divByGcdHelper a b Z gcdPrf
divByGcdGcdOne {a = S a} {b = S b} {c = S c} gcdPrf =
Expand Down
6 changes: 3 additions & 3 deletions libs/contrib/Data/Telescope/Segment.idr
Original file line number Diff line number Diff line change
Expand Up @@ -219,7 +219,7 @@ breakOnto : {0 k,k' : Nat} -> (gamma : Telescope k) -> (pos : Position gamma)
-> {auto 0 ford1 : k' === (finToNat $ finComplement pos) }
-> {default
-- disgusting, sorry
(replace {p = \u => k = finToNat pos + u}
(replace (\u => k = finToNat pos + u)
(sym ford1)
(sym $ finComplementSpec pos))
0 ford2 : (k === ((finToNat pos) + k')) }
Expand Down Expand Up @@ -252,12 +252,12 @@ breakStartEmpty {k} {k' = S k'} {ford1} {ford2} (gamma -. ty) =
v : break {k} {k'} gamma (start gamma) {ford = u}
~=~ Telescope.Nil
v = breakStartEmpty {k} {k'} gamma {ford2 = u}
in replace {p = \z =>
in replace (\z =>
Equal {a = Telescope k} {b = Telescope 0}
(break {k'} {k} gamma (start gamma)
{ford = z})
[]
}
)
(uip u _)
(keep v)

Expand Down
2 changes: 1 addition & 1 deletion libs/contrib/Data/Vect/Properties/Fin.idr
Original file line number Diff line number Diff line change
Expand Up @@ -25,7 +25,7 @@ finNonZero (FS i) = ItIsSucc
||| Inhabitants of `Fin n` witness runtime-irrelevant vectors of length `n` aren't empty
export
finNonEmpty : (0 xs : Vect n a) -> NonZero n -> NonEmpty xs
finNonEmpty xs ItIsSucc = replace {p = NonEmpty} (etaCons xs) IsNonEmpty
finNonEmpty xs ItIsSucc = replace NonEmpty (etaCons xs) IsNonEmpty

||| Cast an index into a runtime-irrelevant `Vect` into the position
||| of the corresponding element
Expand Down
6 changes: 3 additions & 3 deletions libs/contrib/Decidable/Decidable/Extra.idr
Original file line number Diff line number Diff line change
Expand Up @@ -64,9 +64,9 @@ decideTransform :
-> IsDecidable n ts (chain {ts} t r)
decideTransform tDec posDec =
curryAll $ \xs =>
replace {p = id} (chainUncurry (chain t r) Dec xs) $
replace {p = Dec} (chainUncurry r t xs) $
tDec $ replace {p = id} (sym $ chainUncurry r Dec xs) $
replace id (chainUncurry (chain t r) Dec xs) $
replace Dec (chainUncurry r t xs) $
tDec $ replace id (sym $ chainUncurry r Dec xs) $
uncurryAll posDec xs


Expand Down
4 changes: 2 additions & 2 deletions libs/papers/Data/Linear/Inverse.idr
Original file line number Diff line number Diff line change
Expand Up @@ -138,7 +138,7 @@ powMinusAux : (m, n : Nat) -> CmpNat m n ->
Pow a (cast m) -@ Pow a (- cast n) -@ Pow a (cast m - cast n)
powMinusAux m (m + S n) (CmpLT n) pos neg
= let (neg1 # neg2) = powNegativeL m (S n) neg in
powAnnihilate m pos neg1 `seq` replace {p = Pow a} eq neg2
powAnnihilate m pos neg1 `seq` replace (Pow a) eq neg2

where

Expand All @@ -159,7 +159,7 @@ powMinusAux m m CmpEQ pos neg
powAnnihilate (cast m) pos neg
powMinusAux (_ + S m) n (CmpGT m) pos neg
= let (pos1 # pos2) = powPositiveL n (S m) pos in
powAnnihilate n pos1 neg `seq` replace {p = Pow a} eq pos2
powAnnihilate n pos1 neg `seq` replace (Pow a) eq pos2

where

Expand Down
2 changes: 1 addition & 1 deletion libs/papers/Language/IntrinsicTyping/Krivine.idr
Original file line number Diff line number Diff line change
Expand Up @@ -269,7 +269,7 @@ expand {a = Arr a b} {c = ClApp f t} (tr, hored)
= MkPair (step tr)
$ \ arg, red =>
let 0 eq = headReduceClApp (ClApp f t) (\case Lam impossible) arg in
let red = replace {p = Reducible b} (sym eq) (hored arg red) in
let red = replace (Reducible b) (sym eq) (hored arg red) in
expand {c = ClApp (ClApp f t) arg} red

public export
Expand Down
6 changes: 3 additions & 3 deletions libs/papers/Search/Tychonoff/PartI.idr
Original file line number Diff line number Diff line change
Expand Up @@ -52,7 +52,7 @@ record (<->) (a, b : Type) where
map : (a <-> b) -> IsSearchable a -> IsSearchable b
map (MkIso f g _ inv) search p pdec =
let (xa ** prfa) = search (p . f) (\ x => pdec (f x)) in
(f xa ** \ xb, pxb => prfa (g xb) $ replace {p} (sym $ inv xb) pxb)
(f xa ** \ xb, pxb => prfa (g xb) $ replace p (sym $ inv xb) pxb)

interface Searchable (0 a : Type) where
constructor MkSearchable
Expand Down Expand Up @@ -485,7 +485,7 @@ discreteIsUContinuous pdec = MkUC 1 isUContinuous where

isUContinuous : IsUModFor PartI.dc p 1
isUContinuous v w hyp pv with (decEq v w)
_ | Yes eq = replace {p} eq pv
_ | Yes eq = replace p eq pv
_ | No neq = absurd (hyp 0 Oh)

[PROMOTE] Discrete x => CSearchable x PartI.dc => Searchable x where
Expand Down Expand Up @@ -571,7 +571,7 @@ parameters
v0s : Nat -> x; v0s = tail vv0s
in sH .snd v0
$ (sT v0) .snd v0s
$ replace {p} (eta vv0s) pvv0s
$ replace p (eta vv0s) pvv0s

cantorIsCSearchable : Extensionality => IsCSearchable (Nat -> Bool) PartI.dsc
cantorIsCSearchable = csearch @{BYUCONTINUITY}
4 changes: 2 additions & 2 deletions libs/prelude/Builtin.idr
Original file line number Diff line number Diff line change
Expand Up @@ -161,8 +161,8 @@ rewrite__impl p Refl prf = prf
||| Perform substitution in a term according to some equality.
%inline
public export
replace : forall x, y, p . (0 rule : x = y) -> (1 _ : p x) -> p y
replace Refl prf = prf
replace : forall x, y . (0 p : _) -> (0 rule : x = y) -> (1 _ : p x) -> p y
replace _ Refl prf = prf

||| Symmetry of propositional equality.
%inline
Expand Down
2 changes: 1 addition & 1 deletion samples/Proofs.idr
Original file line number Diff line number Diff line change
Expand Up @@ -8,7 +8,7 @@ twoPlusTwo = Refl
-- twoPlusTwoBad = Refl

disjoint : (n : Nat) -> Z = S n -> Void
disjoint n prf = replace {p = disjointTy} prf ()
disjoint n prf = replace disjointTy prf ()
where
disjointTy : Nat -> Type
disjointTy Z = ()
Expand Down
2 changes: 1 addition & 1 deletion src/Core/TT/Var.idr
Original file line number Diff line number Diff line change
Expand Up @@ -283,7 +283,7 @@ insertNVarChiply : SizeOf local ->
insertNVarChiply p v
= rewrite chipsAsListAppend local (n :: outer) in
insertNVar (p <>> zero)
$ replace {p = NVar nm} (chipsAsListAppend local outer) v
$ replace (NVar nm) (chipsAsListAppend local outer) v

export
insertNVarNames : GenWeakenable (NVar name)
Expand Down
2 changes: 1 addition & 1 deletion tests/idris2/data/record021/Issue3501.idr
Original file line number Diff line number Diff line change
Expand Up @@ -2,6 +2,6 @@ import Data.Vect
record Foo (th : Vect n a) where
nIsZero : n === 0
vectIsEmpty : (th ===)
$ replace {p = \ n => Vect n a} (sym nIsZero)
$ replace (\ n => Vect n a) (sym nIsZero)
Copy link
Contributor Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

not sure if I've broken this test, or what it tests

$ Nil

6 changes: 0 additions & 6 deletions tests/idris2/error/perror021/Implicit.idr
Copy link
Contributor Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

deleting this test is likely not the correct action, but I'm unclear what it's testing so waiting for someone to enlighten me

Copy link
Contributor Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

perhaps that a space is expected in {p=\n ...

Copy link
Collaborator

@dunhamsteve dunhamsteve Apr 12, 2025

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

The original commit is #2781 (git has a hard time telling you because the files were moved around). Maybe we could find an example that doesn't use replace. Similar issue, I think it was no space after the =, but the error message itself should point to the correct location rather than saying "couldn't parse declaration". Any implicit pattern match that fails because of a missing space would do. (The "perror" tests are for parse errors.)

This file was deleted.

11 changes: 0 additions & 11 deletions tests/idris2/error/perror021/expected

This file was deleted.

3 changes: 0 additions & 3 deletions tests/idris2/error/perror021/run

This file was deleted.

Loading