mirror of
https://github.com/edwinb/Idris2-boot.git
synced 2024-11-30 22:05:32 +03:00
Merge pull request #195 from ohad/default-totality
Add support for `%default [total | covering | partial]` directive
This commit is contained in:
commit
ff6fd4668b
@ -122,9 +122,6 @@ data Clause : Type where
|
|||||||
MkClause : (env : Env Term vars) ->
|
MkClause : (env : Env Term vars) ->
|
||||||
(lhs : Term vars) -> (rhs : Term vars) -> Clause
|
(lhs : Term vars) -> (rhs : Term vars) -> Clause
|
||||||
|
|
||||||
public export
|
|
||||||
data TotalReq = Total | CoveringOnly | PartialOK
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data DefFlag
|
data DefFlag
|
||||||
= Inline
|
= Inline
|
||||||
@ -142,13 +139,6 @@ data DefFlag
|
|||||||
| BlockedHint -- a hint, but blocked for the moment (so don't use)
|
| BlockedHint -- a hint, but blocked for the moment (so don't use)
|
||||||
| Macro
|
| Macro
|
||||||
|
|
||||||
export
|
|
||||||
Eq TotalReq where
|
|
||||||
(==) Total Total = True
|
|
||||||
(==) CoveringOnly CoveringOnly = True
|
|
||||||
(==) PartialOK PartialOK = True
|
|
||||||
(==) _ _ = False
|
|
||||||
|
|
||||||
export
|
export
|
||||||
Eq DefFlag where
|
Eq DefFlag where
|
||||||
(==) Inline Inline = True
|
(==) Inline Inline = True
|
||||||
@ -1796,6 +1786,13 @@ setUnboundImplicits a
|
|||||||
= do defs <- get Ctxt
|
= do defs <- get Ctxt
|
||||||
put Ctxt (record { options->elabDirectives->unboundImplicits = a } defs)
|
put Ctxt (record { options->elabDirectives->unboundImplicits = a } defs)
|
||||||
|
|
||||||
|
export
|
||||||
|
setDefaultTotalityOption : {auto c : Ref Ctxt Defs} ->
|
||||||
|
TotalReq -> Core ()
|
||||||
|
setDefaultTotalityOption tot
|
||||||
|
= do defs <- get Ctxt
|
||||||
|
put Ctxt (record { options->elabDirectives->totality = tot } defs)
|
||||||
|
|
||||||
export
|
export
|
||||||
isLazyActive : {auto c : Ref Ctxt Defs} ->
|
isLazyActive : {auto c : Ref Ctxt Defs} ->
|
||||||
Core Bool
|
Core Bool
|
||||||
@ -1810,6 +1807,13 @@ isUnboundImplicits
|
|||||||
= do defs <- get Ctxt
|
= do defs <- get Ctxt
|
||||||
pure (unboundImplicits (elabDirectives (options defs)))
|
pure (unboundImplicits (elabDirectives (options defs)))
|
||||||
|
|
||||||
|
export
|
||||||
|
getDefaultTotalityOption : {auto c : Ref Ctxt Defs} ->
|
||||||
|
Core TotalReq
|
||||||
|
getDefaultTotalityOption
|
||||||
|
= do defs <- get Ctxt
|
||||||
|
pure (totality (elabDirectives (options defs)))
|
||||||
|
|
||||||
export
|
export
|
||||||
setPair : {auto c : Ref Ctxt Defs} ->
|
setPair : {auto c : Ref Ctxt Defs} ->
|
||||||
FC -> (pairType : Name) -> (fstn : Name) -> (sndn : Name) ->
|
FC -> (pairType : Name) -> (fstn : Name) -> (sndn : Name) ->
|
||||||
|
@ -2,6 +2,7 @@ module Core.Options
|
|||||||
|
|
||||||
import Core.Core
|
import Core.Core
|
||||||
import Core.Name
|
import Core.Name
|
||||||
|
import Core.TT
|
||||||
import Utils.Binary
|
import Utils.Binary
|
||||||
|
|
||||||
import System.Info
|
import System.Info
|
||||||
@ -84,6 +85,7 @@ record ElabDirectives where
|
|||||||
constructor MkElabDirectives
|
constructor MkElabDirectives
|
||||||
lazyActive : Bool
|
lazyActive : Bool
|
||||||
unboundImplicits : Bool
|
unboundImplicits : Bool
|
||||||
|
totality : TotalReq
|
||||||
|
|
||||||
public export
|
public export
|
||||||
record Session where
|
record Session where
|
||||||
@ -142,7 +144,7 @@ defaultSession : Session
|
|||||||
defaultSession = MkSessionOpts False False False Chez 0 False False
|
defaultSession = MkSessionOpts False False False Chez 0 False False
|
||||||
|
|
||||||
defaultElab : ElabDirectives
|
defaultElab : ElabDirectives
|
||||||
defaultElab = MkElabDirectives True True
|
defaultElab = MkElabDirectives True True PartialOK
|
||||||
|
|
||||||
export
|
export
|
||||||
defaults : Options
|
defaults : Options
|
||||||
|
@ -542,6 +542,16 @@ Ord Visibility where
|
|||||||
compare Public Private = GT
|
compare Public Private = GT
|
||||||
compare Public Export = GT
|
compare Public Export = GT
|
||||||
|
|
||||||
|
public export
|
||||||
|
data TotalReq = Total | CoveringOnly | PartialOK
|
||||||
|
|
||||||
|
export
|
||||||
|
Eq TotalReq where
|
||||||
|
(==) Total Total = True
|
||||||
|
(==) CoveringOnly CoveringOnly = True
|
||||||
|
(==) PartialOK PartialOK = True
|
||||||
|
(==) _ _ = False
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data PartialReason
|
data PartialReason
|
||||||
= NotStrictlyPositive
|
= NotStrictlyPositive
|
||||||
|
@ -566,7 +566,18 @@ mutual
|
|||||||
List Name -> PDecl -> Core (List ImpDecl)
|
List Name -> PDecl -> Core (List ImpDecl)
|
||||||
desugarDecl ps (PClaim fc rig vis fnopts ty)
|
desugarDecl ps (PClaim fc rig vis fnopts ty)
|
||||||
= do opts <- traverse (desugarFnOpt ps) fnopts
|
= do opts <- traverse (desugarFnOpt ps) fnopts
|
||||||
|
opts <- if (isTotalityOption `any` opts)
|
||||||
|
then pure opts
|
||||||
|
else do PartialOK <- getDefaultTotalityOption
|
||||||
|
| tot => pure (Totality tot :: opts)
|
||||||
|
-- We assume PartialOK by default internally
|
||||||
|
pure opts
|
||||||
pure [IClaim fc rig vis opts !(desugarType ps ty)]
|
pure [IClaim fc rig vis opts !(desugarType ps ty)]
|
||||||
|
where
|
||||||
|
isTotalityOption : FnOpt -> Bool
|
||||||
|
isTotalityOption (Totality _) = True
|
||||||
|
isTotalityOption _ = False
|
||||||
|
|
||||||
desugarDecl ps (PDef fc clauses)
|
desugarDecl ps (PDef fc clauses)
|
||||||
-- The clauses won't necessarily all be from the same function, so split
|
-- The clauses won't necessarily all be from the same function, so split
|
||||||
-- after desugaring, by function name, using collectDefs from RawImp
|
-- after desugaring, by function name, using collectDefs from RawImp
|
||||||
@ -727,6 +738,8 @@ mutual
|
|||||||
StartExpr tm => pure [IPragma (\c, nest, env => throw (InternalError "%start not implemented"))] -- TODO!
|
StartExpr tm => pure [IPragma (\c, nest, env => throw (InternalError "%start not implemented"))] -- TODO!
|
||||||
Overloadable n => pure [IPragma (\c, nest, env => setNameFlag fc n Overloadable)]
|
Overloadable n => pure [IPragma (\c, nest, env => setNameFlag fc n Overloadable)]
|
||||||
Extension e => pure [IPragma (\c, nest, env => setExtension e)]
|
Extension e => pure [IPragma (\c, nest, env => setExtension e)]
|
||||||
|
DefaultTotality tot => do setDefaultTotalityOption tot
|
||||||
|
pure []
|
||||||
|
|
||||||
export
|
export
|
||||||
desugar : {auto s : Ref Syn SyntaxInfo} ->
|
desugar : {auto s : Ref Syn SyntaxInfo} ->
|
||||||
|
@ -921,6 +921,15 @@ extension
|
|||||||
= do exactIdent "Borrowing"
|
= do exactIdent "Borrowing"
|
||||||
pure Borrowing
|
pure Borrowing
|
||||||
|
|
||||||
|
totalityOpt : Rule TotalReq
|
||||||
|
totalityOpt
|
||||||
|
= do keyword "partial"
|
||||||
|
pure PartialOK
|
||||||
|
<|> do keyword "total"
|
||||||
|
pure Total
|
||||||
|
<|> do keyword "covering"
|
||||||
|
pure CoveringOnly
|
||||||
|
|
||||||
directive : FileName -> IndentInfo -> Rule Directive
|
directive : FileName -> IndentInfo -> Rule Directive
|
||||||
directive fname indents
|
directive fname indents
|
||||||
= do exactIdent "hide"
|
= do exactIdent "hide"
|
||||||
@ -983,6 +992,10 @@ directive fname indents
|
|||||||
e <- extension
|
e <- extension
|
||||||
atEnd indents
|
atEnd indents
|
||||||
pure (Extension e)
|
pure (Extension e)
|
||||||
|
<|> do exactIdent "default"
|
||||||
|
tot <- totalityOpt
|
||||||
|
atEnd indents
|
||||||
|
pure (DefaultTotality tot)
|
||||||
|
|
||||||
fix : Rule Fixity
|
fix : Rule Fixity
|
||||||
fix
|
fix
|
||||||
@ -1048,16 +1061,11 @@ usingDecls fname indents
|
|||||||
ds <- assert_total (nonEmptyBlock (topDecl fname))
|
ds <- assert_total (nonEmptyBlock (topDecl fname))
|
||||||
end <- location
|
end <- location
|
||||||
pure (PUsing (MkFC fname start end) us (collectDefs (concat ds)))
|
pure (PUsing (MkFC fname start end) us (collectDefs (concat ds)))
|
||||||
|
|
||||||
fnOpt : Rule PFnOpt
|
fnOpt : Rule PFnOpt
|
||||||
fnOpt
|
fnOpt = do x <- totalityOpt
|
||||||
= do keyword "partial"
|
pure $ IFnOpt (Totality x)
|
||||||
pure $ IFnOpt PartialOK
|
|
||||||
<|> do keyword "total"
|
|
||||||
pure $ IFnOpt Total
|
|
||||||
<|> do keyword "covering"
|
|
||||||
pure $ IFnOpt Covering
|
|
||||||
|
|
||||||
fnDirectOpt : FileName -> Rule PFnOpt
|
fnDirectOpt : FileName -> Rule PFnOpt
|
||||||
fnDirectOpt fname
|
fnDirectOpt fname
|
||||||
= do exactIdent "hint"
|
= do exactIdent "hint"
|
||||||
|
@ -161,6 +161,7 @@ mutual
|
|||||||
StartExpr : PTerm -> Directive
|
StartExpr : PTerm -> Directive
|
||||||
Overloadable : Name -> Directive
|
Overloadable : Name -> Directive
|
||||||
Extension : LangExt -> Directive
|
Extension : LangExt -> Directive
|
||||||
|
DefaultTotality : TotalReq -> Directive
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data PField : Type where
|
data PField : Type where
|
||||||
|
@ -65,14 +65,18 @@ visibility
|
|||||||
= visOption
|
= visOption
|
||||||
<|> pure Private
|
<|> pure Private
|
||||||
|
|
||||||
fnOpt : Rule FnOpt
|
totalityOpt : Rule TotalReq
|
||||||
fnOpt
|
totalityOpt
|
||||||
= do keyword "partial"
|
= do keyword "partial"
|
||||||
pure PartialOK
|
pure PartialOK
|
||||||
<|> do keyword "total"
|
<|> do keyword "total"
|
||||||
pure Total
|
pure Total
|
||||||
<|> do keyword "covering"
|
<|> do keyword "covering"
|
||||||
pure Covering
|
pure CoveringOnly
|
||||||
|
|
||||||
|
fnOpt : Rule FnOpt
|
||||||
|
fnOpt = do x <- totalityOpt
|
||||||
|
pure $ Totality x
|
||||||
|
|
||||||
fnDirectOpt : Rule FnOpt
|
fnDirectOpt : Rule FnOpt
|
||||||
fnDirectOpt
|
fnDirectOpt
|
||||||
|
@ -46,12 +46,8 @@ processFnOpt fc ndef (ForeignFn _)
|
|||||||
= setFlag fc ndef Inline -- if externally defined, inline when compiling
|
= setFlag fc ndef Inline -- if externally defined, inline when compiling
|
||||||
processFnOpt fc ndef Invertible
|
processFnOpt fc ndef Invertible
|
||||||
= setFlag fc ndef Invertible
|
= setFlag fc ndef Invertible
|
||||||
processFnOpt fc ndef Total
|
processFnOpt fc ndef (Totality tot)
|
||||||
= setFlag fc ndef (SetTotal Total)
|
= setFlag fc ndef (SetTotal tot)
|
||||||
processFnOpt fc ndef Covering
|
|
||||||
= setFlag fc ndef (SetTotal CoveringOnly)
|
|
||||||
processFnOpt fc ndef PartialOK
|
|
||||||
= setFlag fc ndef (SetTotal PartialOK)
|
|
||||||
processFnOpt fc ndef Macro
|
processFnOpt fc ndef Macro
|
||||||
= setFlag fc ndef Macro
|
= setFlag fc ndef Macro
|
||||||
|
|
||||||
|
@ -187,9 +187,7 @@ mutual
|
|||||||
ForeignFn : List RawImp -> FnOpt
|
ForeignFn : List RawImp -> FnOpt
|
||||||
-- assume safe to cancel arguments in unification
|
-- assume safe to cancel arguments in unification
|
||||||
Invertible : FnOpt
|
Invertible : FnOpt
|
||||||
Total : FnOpt
|
Totality : TotalReq -> FnOpt
|
||||||
Covering : FnOpt
|
|
||||||
PartialOK : FnOpt
|
|
||||||
Macro : FnOpt
|
Macro : FnOpt
|
||||||
|
|
||||||
export
|
export
|
||||||
@ -200,9 +198,9 @@ mutual
|
|||||||
show ExternFn = "%extern"
|
show ExternFn = "%extern"
|
||||||
show (ForeignFn cs) = "%foreign " ++ showSep " " (map show cs)
|
show (ForeignFn cs) = "%foreign " ++ showSep " " (map show cs)
|
||||||
show Invertible = "%invertible"
|
show Invertible = "%invertible"
|
||||||
show Total = "total"
|
show (Totality Total) = "total"
|
||||||
show Covering = "covering"
|
show (Totality CoveringOnly) = "covering"
|
||||||
show PartialOK = "partial"
|
show (Totality PartialOK) = "partial"
|
||||||
show Macro = "%macro"
|
show Macro = "%macro"
|
||||||
|
|
||||||
export
|
export
|
||||||
@ -213,9 +211,7 @@ mutual
|
|||||||
ExternFn == ExternFn = True
|
ExternFn == ExternFn = True
|
||||||
(ForeignFn xs) == (ForeignFn ys) = True -- xs == ys
|
(ForeignFn xs) == (ForeignFn ys) = True -- xs == ys
|
||||||
Invertible == Invertible = True
|
Invertible == Invertible = True
|
||||||
Total == Total = True
|
(Totality tot_lhs) == (Totality tot_rhs) = tot_lhs == tot_rhs
|
||||||
Covering == Covering = True
|
|
||||||
PartialOK == PartialOK = True
|
|
||||||
Macro == Macro = True
|
Macro == Macro = True
|
||||||
_ == _ = False
|
_ == _ = False
|
||||||
|
|
||||||
@ -858,9 +854,9 @@ mutual
|
|||||||
toBuf b ExternFn = tag 3
|
toBuf b ExternFn = tag 3
|
||||||
toBuf b (ForeignFn cs) = do tag 4; toBuf b cs
|
toBuf b (ForeignFn cs) = do tag 4; toBuf b cs
|
||||||
toBuf b Invertible = tag 5
|
toBuf b Invertible = tag 5
|
||||||
toBuf b Total = tag 6
|
toBuf b (Totality Total) = tag 6
|
||||||
toBuf b Covering = tag 7
|
toBuf b (Totality CoveringOnly) = tag 7
|
||||||
toBuf b PartialOK = tag 8
|
toBuf b (Totality PartialOK) = tag 8
|
||||||
toBuf b Macro = tag 9
|
toBuf b Macro = tag 9
|
||||||
|
|
||||||
fromBuf b
|
fromBuf b
|
||||||
@ -871,9 +867,9 @@ mutual
|
|||||||
3 => pure ExternFn
|
3 => pure ExternFn
|
||||||
4 => do cs <- fromBuf b; pure (ForeignFn cs)
|
4 => do cs <- fromBuf b; pure (ForeignFn cs)
|
||||||
5 => pure Invertible
|
5 => pure Invertible
|
||||||
6 => pure Total
|
6 => pure (Totality Total)
|
||||||
7 => pure Covering
|
7 => pure (Totality CoveringOnly)
|
||||||
8 => pure PartialOK
|
8 => pure (Totality PartialOK)
|
||||||
9 => pure Macro
|
9 => pure Macro
|
||||||
_ => corrupt "FnOpt"
|
_ => corrupt "FnOpt"
|
||||||
|
|
||||||
|
Loading…
Reference in New Issue
Block a user