Skip to content

Commit

Permalink
[ new ] Quantity for proof in with-clauses
Browse files Browse the repository at this point in the history
  • Loading branch information
spcfox committed Dec 4, 2024
1 parent ec74792 commit eedfe6d
Show file tree
Hide file tree
Showing 14 changed files with 72 additions and 17 deletions.
2 changes: 2 additions & 0 deletions CHANGELOG_NEXT.md
Original file line number Diff line number Diff line change
Expand Up @@ -71,6 +71,8 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
* Bind expressions in `do` blocks can now have a type ascription.
See [#3327](https://github.com/idris-lang/Idris2/issues/3327).

* Added syntax to specifying quantity for proof in with-clause.

### Compiler changes

* The compiler now differentiates between "package search path" and "package
Expand Down
1 change: 1 addition & 0 deletions CONTRIBUTORS
Original file line number Diff line number Diff line change
Expand Up @@ -81,6 +81,7 @@ Timmy Jose
Tim Süberkrüb
Tom Harley
troiganto
Viktor Yudov
Wen Kokke
Wind Wong
Zoe Stafford
Expand Down
4 changes: 2 additions & 2 deletions libs/base/Language/Reflection/TTImp.idr
Original file line number Diff line number Diff line change
Expand Up @@ -171,7 +171,7 @@ mutual
PatClause : FC -> (lhs : TTImp) -> (rhs : TTImp) -> Clause
WithClause : FC -> (lhs : TTImp) ->
(rig : Count) -> (wval : TTImp) -> -- with'd expression (& quantity)
(prf : Maybe Name) -> -- optional name for the proof
(prf : Maybe (Count, Name)) -> -- optional name for the proof (& quantity)
(flags : List WithFlag) ->
List Clause -> Clause
ImpossibleClause : FC -> (lhs : TTImp) -> Clause
Expand Down Expand Up @@ -605,7 +605,7 @@ mutual
showClause mode (WithClause fc lhs rig wval prf flags cls) -- TODO print flags
= unwords
[ show lhs, "with"
, showCount rig $ maybe id (\ nm => (++ " proof \{show nm}")) prf
, showCount rig $ maybe id (\ nm => (++ " proof \{showCount (fst nm) $ show (snd nm)}")) prf
$ showParens True (show wval)
, "{", joinBy "; " (assert_total $ map (showClause mode) cls), "}"
]
Expand Down
2 changes: 1 addition & 1 deletion src/Idris/Desugar.idr
Original file line number Diff line number Diff line change
Expand Up @@ -926,7 +926,7 @@ mutual
{auto m : Ref MD Metadata} ->
{auto o : Ref ROpts REPLOpts} ->
List Name -> PWithProblem ->
Core (RigCount, RawImp, Maybe Name)
Core (RigCount, RawImp, Maybe (RigCount, Name))
desugarWithProblem ps (MkPWithProblem rig wval mnm)
= (rig,,mnm) <$> desugar AnyExpr ps wval

Expand Down
5 changes: 3 additions & 2 deletions src/Idris/Parser.idr
Original file line number Diff line number Diff line change
Expand Up @@ -1222,8 +1222,9 @@ withProblem fname col indents
= do rig <- multiplicity fname
start <- mustWork $ bounds (decoratedSymbol fname "(")
wval <- bracketedExpr fname start indents
prf <- optional (decoratedKeyword fname "proof"
*> UN . Basic <$> decoratedSimpleBinderName fname)
prf <- optional $ do
decoratedKeyword fname "proof"
pure (!(multiplicity fname), UN $ Basic !(decoratedSimpleBinderName fname))
pure (MkPWithProblem rig wval prf)

mutual
Expand Down
4 changes: 2 additions & 2 deletions src/Idris/REPL.idr
Original file line number Diff line number Diff line change
Expand Up @@ -225,8 +225,8 @@ printClause l i (WithClause _ lhsraw rig wvraw prf flags csraw)
cs <- traverse (printClause l (i + 2)) csraw
pure (relit l ((pack (replicate i ' ')
++ show lhs
++ " with " ++ elimSemi "0 " "1 " (const "") rig ++ "(" ++ show wval ++ ")"
++ maybe "" (\ nm => " proof " ++ show nm) prf
++ " with " ++ showCount rig ++ "(" ++ show wval ++ ")"
++ maybe "" (\ nm => " proof " ++ showCount (fst nm) ++ show (snd nm)) prf
++ "\n"))
++ showSep "\n" cs)
printClause l i (ImpossibleClause _ lhsraw)
Expand Down
2 changes: 1 addition & 1 deletion src/Idris/Syntax.idr
Original file line number Diff line number Diff line change
Expand Up @@ -342,7 +342,7 @@ mutual
constructor MkPWithProblem
withRigCount : RigCount
withRigValue : PTerm' nm
withRigProof : Maybe Name -- This ought to be a `Basic` username
withRigProof : Maybe (RigCount, Name) -- This ought to be a `Basic` username

public export
PClause : Type
Expand Down
4 changes: 3 additions & 1 deletion src/TTImp/Parser.idr
Original file line number Diff line number Diff line change
Expand Up @@ -522,7 +522,9 @@ mutual
symbol "("
wval <- expr fname indents
symbol ")"
prf <- optional (keyword "proof" *> name)
prf <- optional $ do
keyword "proof"
pure (!(getMult !multiplicity), !name)
ws <- nonEmptyBlock (clause (S withArgs) fname)
end <- location
let fc = MkFC fname start end
Expand Down
8 changes: 4 additions & 4 deletions src/TTImp/ProcessDef.idr
Original file line number Diff line number Diff line change
Expand Up @@ -600,7 +600,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env
Just _ =>
let fc = emptyFC in
let refl = IVar fc (NS builtinNS (UN $ Basic "Refl")) in
[(mprf, INamedApp fc refl (UN $ Basic "x") wval_raw)]
[(map snd mprf, INamedApp fc refl (UN $ Basic "x") wval_raw)]

let rhs_in = gapply (IVar vfc wname)
$ map (\ nm => (Nothing, IVar vfc nm)) envns
Expand Down Expand Up @@ -633,7 +633,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env
mkExplicit (b :: env) = b :: mkExplicit env

bindWithArgs :
(rig : RigCount) -> (wvalTy : Term xs) -> Maybe (Name, Term xs) ->
(rig : RigCount) -> (wvalTy : Term xs) -> Maybe ((RigCount, Name), Term xs) ->
(wvalEnv : Env Term xs) ->
Core (ext : List Name
** ( Env Term (ext ++ xs)
Expand All @@ -657,7 +657,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env

in pure (wargs ** (scenv, var, binder))

bindWithArgs {xs} rig wvalTy (Just (name, wval)) wvalEnv = do
bindWithArgs {xs} rig wvalTy (Just ((rigPrf, name), wval)) wvalEnv = do
defs <- get Ctxt

let eqName = NS builtinNS (UN $ Basic "Equal")
Expand Down Expand Up @@ -689,7 +689,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env

binder : Term (wargs ++ xs) -> Term xs
:= \ t => Bind vfc wargn (Pi vfc rig Explicit wvalTy)
$ Bind vfc name (Pi vfc rig Implicit eqTy) t
$ Bind vfc name (Pi vfc rigPrf Implicit eqTy) t

pure (wargs ** (scenv, var, binder))

Expand Down
8 changes: 4 additions & 4 deletions src/TTImp/TTImp.idr
Original file line number Diff line number Diff line change
Expand Up @@ -427,7 +427,7 @@ mutual
PatClause : FC -> (lhs : RawImp' nm) -> (rhs : RawImp' nm) -> ImpClause' nm
WithClause : FC -> (lhs : RawImp' nm) ->
(rig : RigCount) -> (wval : RawImp' nm) -> -- with'd expression (& quantity)
(prf : Maybe Name) -> -- optional name for the proof
(prf : Maybe (RigCount, Name)) -> -- optional name for the proof
(flags : List WithFlag) ->
List (ImpClause' nm) -> ImpClause' nm
ImpossibleClause : FC -> (lhs : RawImp' nm) -> ImpClause' nm
Expand All @@ -441,8 +441,8 @@ mutual
= show lhs ++ " = " ++ show rhs
show (WithClause fc lhs rig wval prf flags block)
= show lhs
++ " with (" ++ show rig ++ " " ++ show wval ++ ")"
++ maybe "" (\ nm => " proof " ++ show nm) prf
++ " with " ++ showCount rig ++ "(" ++ show wval ++ ")"
++ maybe "" (\ nm => " proof " ++ showCount (fst nm) ++ show (snd nm)) prf
++ "\n\t" ++ show block
show (ImpossibleClause fc lhs)
= show lhs ++ " impossible"
Expand Down Expand Up @@ -518,7 +518,7 @@ mutual


export
mkWithClause : FC -> RawImp' nm -> List1 (RigCount, RawImp' nm, Maybe Name) ->
mkWithClause : FC -> RawImp' nm -> List1 (RigCount, RawImp' nm, Maybe (RigCount, Name)) ->
List WithFlag -> List (ImpClause' nm) -> ImpClause' nm
mkWithClause fc lhs ((rig, wval, prf) ::: []) flags cls
= WithClause fc lhs rig wval prf flags cls
Expand Down
20 changes: 20 additions & 0 deletions tests/idris2/with/with012/WithProof0.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,20 @@
module WithProof0

%default total

%hide List.filter

filter : (p : a -> Bool) -> (xs : List a) -> List a
filter p [] = []
filter p (x :: xs) with (p x)
filter p (x :: xs) | False = filter p xs
filter p (x :: xs) | True = x :: filter p xs


filterSquared : (p : a -> Bool) -> (xs : List a) ->
filter p (filter p xs) === filter p xs
filterSquared p [] = Refl
filterSquared p (x :: xs) with (p x) proof 0 eq
filterSquared p (x :: xs) | False = filterSquared p xs -- easy
filterSquared p (x :: xs) | True
= rewrite eq in cong (x ::) (filterSquared p xs)
23 changes: 23 additions & 0 deletions tests/idris2/with/with012/WithProof1.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,23 @@
module WithProof1

%default total

%hide List.filter

filter : (p : a -> Bool) -> (xs : List a) -> List a
filter p [] = []
filter p (x :: xs) with (p x)
filter p (x :: xs) | False = filter p xs
filter p (x :: xs) | True = x :: filter p xs

consumeEq : (1 _ : x === y) -> ()
consumeEq Refl = ()

filterSquared : (p : a -> Bool) -> (xs : List a) ->
filter p (filter p xs) === filter p xs
filterSquared p [] = Refl
filterSquared p (x :: xs) with (p x) proof 1 eq
filterSquared p (x :: xs) | False = case consumeEq eq of
() => filterSquared p xs
filterSquared p (x :: xs) | True = case consumeEq eq of
() => rewrite eq in cong (x ::) (filterSquared p xs)
2 changes: 2 additions & 0 deletions tests/idris2/with/with012/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,2 @@
1/1: Building WithProof0 (WithProof0.idr)
1/1: Building WithProof1 (WithProof1.idr)
4 changes: 4 additions & 0 deletions tests/idris2/with/with012/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,4 @@
. ../../../testutils.sh

check WithProof0.idr
check WithProof1.idr

0 comments on commit eedfe6d

Please sign in to comment.