Skip to content

Commit

Permalink
Merge branch 'maint-2.4.2'
Browse files Browse the repository at this point in the history
  • Loading branch information
andreasabel committed Oct 28, 2015
2 parents a8b6097 + bb8d1ac commit 63674c9
Show file tree
Hide file tree
Showing 5 changed files with 142 additions and 9 deletions.
15 changes: 15 additions & 0 deletions src/full/Agda/Termination/Inlining.hs
Original file line number Diff line number Diff line change
Expand Up @@ -87,15 +87,27 @@ inlineWithClauses f cl = inTopContext $ do
-- The clause body is a with-function call @wf args@.
-- @f@ is the function the with-function belongs to.
let args = fromMaybe __IMPOSSIBLE__ . allApplyElims $ els

reportSDoc "term.with.inline" 70 $ sep
[ text "Found with (raw):", nest 2 $ text $ show cl ]
reportSDoc "term.with.inline" 20 $ sep
[ text "Found with:", nest 2 $ prettyTCM $ QNamed f cl ]

t <- defType <$> getConstInfo wf
cs1 <- withExprClauses cl t args

reportSDoc "term.with.inline" 70 $ vcat $
text "withExprClauses (raw)" : map (nest 2 . text . show) cs1
reportSDoc "term.with.inline" 20 $ vcat $
text "withExprClauses" : map (nest 2 . prettyTCM . QNamed f) cs1

cs2 <- inlinedClauses f cl t wf

reportSDoc "term.with.inline" 70 $ vcat $
text "inlinedClauses (raw)" : map (nest 2 . text . show) cs2
reportSDoc "term.with.inline" 20 $ vcat $
text "inlinedClauses" : map (nest 2 . prettyTCM . QNamed f) cs2

return $ cs1 ++ cs2
_ -> noInline

Expand Down Expand Up @@ -153,8 +165,11 @@ inline f pcl t wf wcl = inTopContext $ addCtxTel (clauseTel wcl) $ do
-- of the arguments to the parent function. Fortunately we have already
-- figured out how to turn an application of the with-function into an
-- application of the parent function in the display form.
reportSDoc "term.with.inline" 70 $ text "inlining (raw) =" <+> text (show wcl)
let vs = clauseArgs wcl
Just disp <- displayForm wf vs
reportSDoc "term.with.inline" 70 $ text "display form (raw) =" <+> text (show disp)
reportSDoc "term.with.inline" 40 $ text "display form =" <+> prettyTCM disp
(pats, perm) <- dispToPats disp

-- Now we need to sort out the right hand side. We have
Expand Down
2 changes: 1 addition & 1 deletion src/full/Agda/TypeChecking/Pretty.hs
Original file line number Diff line number Diff line change
Expand Up @@ -347,7 +347,7 @@ showPat' :: (a -> TCM Doc) -> Pattern' a -> TCM Doc
showPat' showVar = showPat
where
showPat (VarP x) = showVar x
showPat (DotP t) = text $ ".(" ++ show t ++ ")"
showPat (DotP t) = text $ ".(" ++ P.prettyShow t ++ ")"
showPat (ConP c i ps) = (if b then braces else parens) $ prTy $
prettyTCM c <+> fsep (map (showPat . namedArg) ps)
where
Expand Down
23 changes: 15 additions & 8 deletions src/full/Agda/TypeChecking/With.hs
Original file line number Diff line number Diff line change
Expand Up @@ -35,6 +35,7 @@ import Agda.Utils.Functor
import Agda.Utils.List
import Agda.Utils.Monad
import Agda.Utils.Permutation
import Agda.Utils.Pretty (prettyShow)
import Agda.Utils.Size

#include "undefined.h"
Expand Down Expand Up @@ -398,7 +399,8 @@ withDisplayForm f aux delta1 delta2 n qs perm@(Perm m _) lhsPerm = do
-- Build a substitution to replace the parent pattern vars
-- by the pattern vars of the with-function.
(ys0, ys1) = splitAt (size delta1) $ permute perm $ downFrom m
ys = reverse $ map Just ys0 ++ replicate n Nothing ++ map Just ys1
ys = reverse (map Just ys0 ++ replicate n Nothing ++ map Just ys1)
++ map (Just . (m +)) [0..top-1]
rho = sub top ys wild
tqs = applySubst rho tqs0
-- Build the arguments to the with function.
Expand All @@ -424,22 +426,27 @@ withDisplayForm f aux delta1 delta2 n qs perm@(Perm m _) lhsPerm = do
, text "n =" <+> text (show n)
, text "perm =" <+> text (show perm)
, text "top =" <+> do addFullCtx $ prettyTCM topArgs
, text "qs =" <+> text (show qs)
, text "dt =" <+> do addFullCtx $ prettyTCM dt
, text "ys =" <+> text (show ys)
, text "raw =" <+> text (show display)
, text "qs =" <+> sep (map (prettyTCM . namedArg) qs)
, text "qsToTm =" <+> prettyTCM tqs0 -- ctx would be permuted form of delta1 ++ delta2
, text "sub qs =" <+> prettyTCM tqs
, text "ys =" <+> text (show ys)
, text "rho =" <+> text (prettyShow rho)
, text "qs[rho]=" <+> do addFullCtx $ prettyTCM tqs
, text "dt =" <+> do addFullCtx $ prettyTCM dt
]
]
reportSDoc "tc.with.display" 70 $ nest 2 $ vcat
[ text "raw =" <+> text (show display)
]

return display
where
-- Ulf, 2014-02-19: We need to rename the module parameters as well! (issue1035)
sub top ys wild = map term [0 .. m - 1] ++# raiseS (length qs)
-- sub top ys wild = map term [0 .. m - 1] ++# raiseS (length qs)
-- Andreas, 2015-10-28: Yes, but properly! (Issue 1407)
sub top ys wild = parallelS $ map term [0 .. m + top - 1]
where
term i = maybe wild var $ findIndex (Just i ==) ys
-- OLD
-- -- OLD
-- sub top rho wild = parallelS $ map term [0 .. m - 1] ++ topTerms
-- where
-- -- Ulf, 2014-02-19: We need to rename the module parameters as well! (issue1035)
Expand Down
24 changes: 24 additions & 0 deletions test/succeed/Issue1035.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,24 @@
-- Andreas, 2015-10-28 adapted from test case by
-- Ulf, 2014-02-06, shrunk from Wolfram Kahl's report

open import Common.Equality
open import Common.Unit
open import Common.Nat

postulate
prop : x x ≡ unit

record StrictTotalOrder : Set where
field compare : Unit
open StrictTotalOrder

module M (Key : StrictTotalOrder) where

postulate
intersection′-₁ : x x ≡ compare Key

to-∈-intersection′ : Nat Unit Unit Set
to-∈-intersection′ zero x h = Unit
to-∈-intersection′ (suc n) x h with intersection′-₁ x
to-∈-intersection′ (suc n) ._ h | refl with prop h
to-∈-intersection′ (suc n) ._ ._ | refl | refl = to-∈-intersection′ n unit unit
87 changes: 87 additions & 0 deletions test/succeed/Issue1407.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,87 @@
-- Andreas, 2015-10-28 Issue 1407
-- (Issue 1023 raised its ugly head again!)

{-# OPTIONS -v term.with.inline:30 #-} --KEEP! Triggered bug on agda-2.4.2.5.
-- {-# OPTIONS -v term.with.inline:70 #-}
-- {-# OPTIONS -v tc.with.display:30 #-}

record _×_ (A B : Set) : Set where
constructor pair
field fst : A
snd : B
open _×_

data Unit : Set where
unit : Unit

data List (A : Set) : Set where
_∷_ : A List A List A

map : A B (A B) List A List B
map A B f (x ∷ xs) = f x ∷ map A B f xs

record Setoid : Set₁ where
field Carrier : Set
open Setoid

module Assoc (Key : Setoid) (_ : Set) where -- FAILS
-- module Assoc (_ : Set) (Key : Setoid) where -- WORKS

data _∈_ (x : Carrier Key) : List (Carrier Key) Set where
here : xs Unit x ∈ xs
there : {xs} x ∈ xs

postulate
foo : Carrier Key Carrier Key Unit

lem : (p : Carrier Key × Unit)
(ps : List (Carrier Key × Unit))
fst p ∈ map (Carrier Key × Unit) (Carrier Key) fst ps
Set
lem (pair k v) ((pair k' v') ∷ ps) (here .(k' ∷ map (Carrier Key × Unit) (Carrier Key) fst ps) u) with foo k k'
... | unit = Unit
lem p (_ ∷ ps) there = lem p ps there -- to force termination checking

-- 7 6 5 4 3 2 1 7 0
-- lem Key _ (pair k v) (pair k' v') ∷ ps) (here .(...Key...) u)
-- 8 7 6 5 4 3 2 1 0
-- lem-with Key _ k k' (w = unit) v v' ps u = Unit
{-
inlinedClauses (raw)
QNamed {qname = Issue1407.Assoc.lem, qnamed = Clause
{ clauseRange = /home/abel/agda/test/bugs/Issue1407.agda:40,3-13
, clauseTel =
ExtendTel (Def Issue1407.Setoid []) (Abs "Key"
ExtendTel (Sort (Type (Max []))) (Abs "_"
ExtendTel (Var 1 [Proj Issue1407.Setoid.Carrier]}) (Abs "k"
ExtendTel (Var 2 [Proj Issue1407.Setoid.Carrier]}) (Abs "k'"
ExtendTel (Def Issue1407.Unit []}) (Abs "v"
ExtendTel (Def Issue1407.Unit []}) (Abs "v'"
ExtendTel (Def Issue1407.List [Apply []r(Def Issue1407._×_ [Apply []r(Var 5 [Proj Issue1407.Setoid.Carrier]),Apply []r(Def Issue1407.Unit [])])]}) (Abs "ps"
ExtendTel (Def Issue1407.Unit []}) (Abs "u"
EmptyTel))))))))
, clausePerm = x0,x1,x2,x3,x4,x5,x6,x7,x8 -> x0,x1,x2,x4,x3,x5,x6,x8
, namedClausePats =
[[]r(VarP "Key")
,[]r(VarP "x")
,[]r(ConP Issue1407._×_._,_(inductive)[Issue1407._×_.fst,Issue1407._×_.snd] (ConPatternInfo {conPRecord = Nothing, conPType = Nothing}) [[]r(VarP "k"),[]r(VarP "v")])
,[]r(ConP Issue1407.List._∷_(inductive)[] (ConPatternInfo {conPRecord = Nothing, conPType = Nothing}) [[]r(ConP Issue1407._×_._,_(inductive)[Issue1407._×_.fst,Issue1407._×_.snd] (ConPatternInfo {conPRecord = Nothing, conPType = Nothing}) [[]r(VarP "k'"),[]r(VarP "v'")]),[]r(VarP "ps")])
,[]r(ConP Issue1407.Assoc._∈_.here(inductive)[] (ConPatternInfo {conPRecord = Nothing, conPType = Nothing})
[[]r{DotP (Con Issue1407.List._∷_(inductive)[]
[[]r(Var 4 [])
,[]r(Def Issue1407.map
[Apply []r{Def Issue1407._×_ [Apply []r(Def Issue1407.Assoc._.Carrier [Apply []r(Con Issue1407.Unit.unit(inductive)[] []),Apply []r(Var 3 [])])
,Apply []r(Def Issue1407.Unit [])]}
,Apply []r{Def Issue1407.Assoc._.Carrier [Apply []r(Con Issue1407.Unit.unit(inductive)[] [])
,Apply []r(Var 3 [])]}
,Apply []r(Lam (ArgInfo {argInfoHiding = NotHidden, argInfoRelevance = Relevant, argInfoColors = []}) (Abs "r" Var 0 [Proj Issue1407._×_.fst]))
,Apply []r(Var 1 [])
])
])}
,[]r(VarP "u")
])
,[]r(ConP Issue1407.Unit.unit(inductive)[] (ConPatternInfo {conPRecord = Nothing, conPType = Nothing}) [])
]
, clauseBody = Bind (Abs "h8" Bind (Abs "h7" Bind (Abs "h6" Bind (Abs "h5" Bind (Abs "h4" Bind (Abs "h3" Bind (Abs "h2" Bind (Abs "h1" Bind (Abs "h0" Body (Def Issue1407.Unit [])))))))))), clauseType = Nothing}}
-}

0 comments on commit 63674c9

Please sign in to comment.