Tidy up the handling of wild-card binders, and make Lint check it
[ghc-hetmet.git] / compiler / coreSyn / CoreLint.lhs
1
2 %
3 % (c) The University of Glasgow 2006
4 % (c) The GRASP/AQUA Project, Glasgow University, 1993-1998
5 %
6
7 A ``lint'' pass to check for Core correctness
8
9 \begin{code}
10 module CoreLint ( lintCoreBindings, lintUnfolding ) where
11
12 #include "HsVersions.h"
13
14 import Demand
15 import CoreSyn
16 import CoreFVs
17 import CoreUtils
18 import Bag
19 import Literal
20 import DataCon
21 import TysWiredIn
22 import Var
23 import VarEnv
24 import VarSet
25 import Name
26 import Id
27 import PprCore
28 import ErrUtils
29 import SrcLoc
30 import Type
31 import TypeRep
32 import Coercion
33 import TyCon
34 import Class
35 import BasicTypes
36 import StaticFlags
37 import ListSetOps
38 import PrelNames
39 import Outputable
40 import FastString
41 import Util
42 import Control.Monad
43 import Data.Maybe
44 \end{code}
45
46 %************************************************************************
47 %*                                                                      *
48 \subsection[lintCoreBindings]{@lintCoreBindings@: Top-level interface}
49 %*                                                                      *
50 %************************************************************************
51
52 Checks that a set of core bindings is well-formed.  The PprStyle and String
53 just control what we print in the event of an error.  The Bool value
54 indicates whether we have done any specialisation yet (in which case we do
55 some extra checks).
56
57 We check for
58         (a) type errors
59         (b) Out-of-scope type variables
60         (c) Out-of-scope local variables
61         (d) Ill-kinded types
62
63 If we have done specialisation the we check that there are
64         (a) No top-level bindings of primitive (unboxed type)
65
66 Outstanding issues:
67
68     --
69     -- Things are *not* OK if:
70     --
71     --  * Unsaturated type app before specialisation has been done;
72     --
73     --  * Oversaturated type app after specialisation (eta reduction
74     --   may well be happening...);
75
76
77 Note [Linting type lets]
78 ~~~~~~~~~~~~~~~~~~~~~~~~
79 In the desugarer, it's very very convenient to be able to say (in effect)
80         let a = Type Int in <body>
81 That is, use a type let.   See Note [Type let] in CoreSyn.
82
83 However, when linting <body> we need to remember that a=Int, else we might
84 reject a correct program.  So we carry a type substitution (in this example 
85 [a -> Int]) and apply this substitution before comparing types.  The functin
86         lintInTy :: Type -> LintM Type
87 returns a substituted type; that's the only reason it returns anything.
88
89 When we encounter a binder (like x::a) we must apply the substitution
90 to the type of the binding variable.  lintBinders does this.
91
92 For Ids, the type-substituted Id is added to the in_scope set (which 
93 itself is part of the TvSubst we are carrying down), and when we
94 find an occurence of an Id, we fetch it from the in-scope set.
95
96
97 \begin{code}
98 lintCoreBindings :: [CoreBind] -> (Bag Message, Bag Message)
99 --   Returns (warnings, errors)
100 lintCoreBindings binds
101   = initL $ 
102     addLoc TopLevelBindings $
103     addInScopeVars binders  $
104         -- Put all the top-level binders in scope at the start
105         -- This is because transformation rules can bring something
106         -- into use 'unexpectedly'
107     do { checkL (null dups) (dupVars dups)
108        ; checkL (null ext_dups) (dupExtVars ext_dups)
109        ; mapM lint_bind binds }
110   where
111     binders = bindersOfBinds binds
112     (_, dups) = removeDups compare binders
113
114     -- dups_ext checks for names with different uniques
115     -- but but the same External name M.n.  We don't
116     -- allow this at top level:
117     --    M.n{r3}  = ...
118     --    M.n{r29} = ...
119     -- becuase they both get the same linker symbol
120     ext_dups = snd (removeDups ord_ext (map Var.varName binders))
121     ord_ext n1 n2 | Just m1 <- nameModule_maybe n1
122                   , Just m2 <- nameModule_maybe n2
123                   = compare (m1, nameOccName n1) (m2, nameOccName n2)
124                   | otherwise = LT
125
126     lint_bind (Rec prs)         = mapM_ (lintSingleBinding TopLevel Recursive) prs
127     lint_bind (NonRec bndr rhs) = lintSingleBinding TopLevel NonRecursive (bndr,rhs)
128 \end{code}
129
130 %************************************************************************
131 %*                                                                      *
132 \subsection[lintUnfolding]{lintUnfolding}
133 %*                                                                      *
134 %************************************************************************
135
136 We use this to check all unfoldings that come in from interfaces
137 (it is very painful to catch errors otherwise):
138
139 \begin{code}
140 lintUnfolding :: SrcLoc
141               -> [Var]          -- Treat these as in scope
142               -> CoreExpr
143               -> Maybe Message  -- Nothing => OK
144
145 lintUnfolding locn vars expr
146   | isEmptyBag errs = Nothing
147   | otherwise       = Just (pprMessageBag errs)
148   where
149     (_warns, errs) = initL (addLoc (ImportedUnfolding locn) $
150                             addInScopeVars vars            $
151                             lintCoreExpr expr)
152 \end{code}
153
154 %************************************************************************
155 %*                                                                      *
156 \subsection[lintCoreBinding]{lintCoreBinding}
157 %*                                                                      *
158 %************************************************************************
159
160 Check a core binding, returning the list of variables bound.
161
162 \begin{code}
163 lintSingleBinding :: TopLevelFlag -> RecFlag -> (Id, CoreExpr) -> LintM ()
164 lintSingleBinding top_lvl_flag rec_flag (binder,rhs)
165   = addLoc (RhsOf binder) $
166          -- Check the rhs 
167     do { ty <- lintCoreExpr rhs 
168        ; lintBinder binder -- Check match to RHS type
169        ; binder_ty <- applySubst binder_ty
170        ; checkTys binder_ty ty (mkRhsMsg binder ty)
171         -- Check (not isUnLiftedType) (also checks for bogus unboxed tuples)
172        ; checkL (not (isUnLiftedType binder_ty)
173             || (isNonRec rec_flag && exprOkForSpeculation rhs))
174            (mkRhsPrimMsg binder rhs)
175         -- Check that if the binder is top-level or recursive, it's not demanded
176        ; checkL (not (isStrictId binder)
177             || (isNonRec rec_flag && not (isTopLevel top_lvl_flag)))
178            (mkStrictMsg binder)
179         -- Check whether binder's specialisations contain any out-of-scope variables
180        ; mapM_ (checkBndrIdInScope binder) bndr_vars 
181
182        ; when (isNonRuleLoopBreaker (idOccInfo binder) && isInlinePragma (idInlinePragma binder))
183               (addWarnL (ptext (sLit "INLINE binder is (non-rule) loop breaker:") <+> ppr binder))
184               -- Only non-rule loop breakers inhibit inlining
185
186       -- Check whether arity and demand type are consistent (only if demand analysis
187       -- already happened)
188        ; checkL (case maybeDmdTy of
189                   Just (StrictSig dmd_ty) -> idArity binder >= dmdTypeDepth dmd_ty || exprIsTrivial rhs
190                   Nothing -> True)
191            (mkArityMsg binder) }
192           
193         -- We should check the unfolding, if any, but this is tricky because
194         -- the unfolding is a SimplifiableCoreExpr. Give up for now.
195    where
196     binder_ty                  = idType binder
197     maybeDmdTy                 = idStrictness_maybe binder
198     bndr_vars                  = varSetElems (idFreeVars binder)
199     lintBinder var | isId var  = lintIdBndr var $ \_ -> (return ())
200                    | otherwise = return ()
201 \end{code}
202
203 %************************************************************************
204 %*                                                                      *
205 \subsection[lintCoreExpr]{lintCoreExpr}
206 %*                                                                      *
207 %************************************************************************
208
209 \begin{code}
210 type InType  = Type     -- Substitution not yet applied
211 type InVar   = Var
212 type InTyVar = TyVar
213
214 type OutType  = Type    -- Substitution has been applied to this
215 type OutVar   = Var
216 type OutTyVar = TyVar
217 type OutCoVar = CoVar
218
219 lintCoreExpr :: CoreExpr -> LintM OutType
220 -- The returned type has the substitution from the monad 
221 -- already applied to it:
222 --      lintCoreExpr e subst = exprType (subst e)
223 --
224 -- The returned "type" can be a kind, if the expression is (Type ty)
225
226 lintCoreExpr (Var var)
227   = do  { checkL (not (var == oneTupleDataConId))
228                  (ptext (sLit "Illegal one-tuple"))
229
230         ; checkL (not (var `hasKey` wildCardKey))
231                  (ptext (sLit "Occurence of a wild-card binder") <+> ppr var)
232                  -- See Note [WildCard binders] in SimplEnv
233
234         ; checkDeadIdOcc var
235         ; var' <- lookupIdInScope var
236         ; return (idType var') }
237
238 lintCoreExpr (Lit lit)
239   = return (literalType lit)
240
241 lintCoreExpr (Cast expr co)
242   = do { expr_ty <- lintCoreExpr expr
243        ; co' <- applySubst co
244        ; (from_ty, to_ty) <- lintCoercion co'
245        ; checkTys from_ty expr_ty (mkCastErr from_ty expr_ty)
246        ; return to_ty }
247
248 lintCoreExpr (Note _ expr)
249   = lintCoreExpr expr
250
251 lintCoreExpr (Let (NonRec tv (Type ty)) body)
252   | isTyVar tv
253   =     -- See Note [Linting type lets]
254     do  { ty' <- addLoc (RhsOf tv) $ lintInTy ty
255         ; lintTyBndr tv              $ \ tv' -> 
256           addLoc (BodyOfLetRec [tv]) $ 
257           extendSubstL tv' ty'       $ do
258         { checkKinds tv' ty'              
259                 -- Now extend the substitution so we 
260                 -- take advantage of it in the body
261         ; lintCoreExpr body } }
262
263   | isCoVar tv
264   = do { co <- applySubst ty
265        ; (s1,s2) <- addLoc (RhsOf tv) $ lintCoercion co
266        ; lintTyBndr tv  $ \ tv' -> 
267          addLoc (BodyOfLetRec [tv]) $ do
268        { let (t1,t2) = coVarKind tv'
269        ; checkTys s1 t1 (mkTyVarLetErr tv ty)
270        ; checkTys s2 t2 (mkTyVarLetErr tv ty)
271        ; lintCoreExpr body } }
272
273   | otherwise
274   = failWithL (mkTyVarLetErr tv ty)     -- Not quite accurate
275
276 lintCoreExpr (Let (NonRec bndr rhs) body)
277   = do  { lintSingleBinding NotTopLevel NonRecursive (bndr,rhs)
278         ; addLoc (BodyOfLetRec [bndr])
279                  (lintAndScopeId bndr $ \_ -> (lintCoreExpr body)) }
280
281 lintCoreExpr (Let (Rec pairs) body) 
282   = lintAndScopeIds bndrs       $ \_ ->
283     do  { checkL (null dups) (dupVars dups)
284         ; mapM_ (lintSingleBinding NotTopLevel Recursive) pairs 
285         ; addLoc (BodyOfLetRec bndrs) (lintCoreExpr body) }
286   where
287     bndrs = map fst pairs
288     (_, dups) = removeDups compare bndrs
289
290 lintCoreExpr e@(App fun arg)
291   = do  { fun_ty <- lintCoreExpr fun
292         ; addLoc (AnExpr e) $
293           lintCoreArg fun_ty arg }
294
295 lintCoreExpr (Lam var expr)
296   = addLoc (LambdaBodyOf var) $
297     lintBinders [var] $ \ vars' ->
298     do { let [var'] = vars'  
299        ; body_ty <- lintCoreExpr expr
300        ; if isId var' then 
301              return (mkFunTy (idType var') body_ty) 
302          else
303              return (mkForAllTy var' body_ty)
304        }
305         -- The applySubst is needed to apply the subst to var
306
307 lintCoreExpr e@(Case scrut var alt_ty alts) =
308        -- Check the scrutinee
309   do { scrut_ty <- lintCoreExpr scrut
310      ; alt_ty   <- lintInTy alt_ty  
311      ; var_ty   <- lintInTy (idType var)        
312
313      ; let mb_tc_app = splitTyConApp_maybe (idType var)
314      ; case mb_tc_app of 
315          Just (tycon, _)
316               | debugIsOn &&
317                 isAlgTyCon tycon && 
318                 not (isFamilyTyCon tycon || isAbstractTyCon tycon) &&
319                 null (tyConDataCons tycon) -> 
320                   pprTrace "Lint warning: case binder's type has no constructors" (ppr var <+> ppr (idType var))
321                         -- This can legitimately happen for type families
322                       $ return ()
323          _otherwise -> return ()
324
325         -- Don't use lintIdBndr on var, because unboxed tuple is legitimate
326
327      ; subst <- getTvSubst 
328      ; checkTys var_ty scrut_ty (mkScrutMsg var var_ty scrut_ty subst)
329
330      -- If the binder is an unboxed tuple type, don't put it in scope
331      ; let scope = if (isUnboxedTupleType (idType var)) then 
332                        pass_var 
333                    else lintAndScopeId var
334      ; scope $ \_ ->
335        do { -- Check the alternatives
336             mapM_ (lintCoreAlt scrut_ty alt_ty) alts
337           ; checkCaseAlts e scrut_ty alts
338           ; return alt_ty } }
339   where
340     pass_var f = f var
341
342 lintCoreExpr (Type ty)
343   = do { ty' <- lintInTy ty
344        ; return (typeKind ty') }
345 \end{code}
346
347 %************************************************************************
348 %*                                                                      *
349 \subsection[lintCoreArgs]{lintCoreArgs}
350 %*                                                                      *
351 %************************************************************************
352
353 The basic version of these functions checks that the argument is a
354 subtype of the required type, as one would expect.
355
356 \begin{code}
357 lintCoreArg  :: OutType -> CoreArg -> LintM OutType
358 lintCoreArg fun_ty (Type arg_ty)
359   = do  { arg_ty' <- applySubst arg_ty
360         ; lintTyApp fun_ty arg_ty' }
361
362 lintCoreArg fun_ty arg
363  = do { arg_ty <- lintCoreExpr arg
364       ; lintValApp arg fun_ty arg_ty }
365
366 -----------------
367 lintAltBinders :: OutType     -- Scrutinee type
368                -> OutType     -- Constructor type
369                -> [OutVar]    -- Binders
370                -> LintM ()
371 lintAltBinders scrut_ty con_ty [] 
372   = checkTys con_ty scrut_ty (mkBadPatMsg con_ty scrut_ty) 
373 lintAltBinders scrut_ty con_ty (bndr:bndrs)
374   | isTyCoVar bndr
375   = do { con_ty' <- lintTyApp con_ty (mkTyVarTy bndr)
376        ; lintAltBinders scrut_ty con_ty' bndrs }
377   | otherwise
378   = do { con_ty' <- lintValApp (Var bndr) con_ty (idType bndr)
379        ; lintAltBinders scrut_ty con_ty' bndrs } 
380
381 -----------------
382 lintTyApp :: OutType -> OutType -> LintM OutType
383 lintTyApp fun_ty arg_ty
384   | Just (tyvar,body_ty) <- splitForAllTy_maybe fun_ty
385   = do  { checkKinds tyvar arg_ty
386         ; if isCoVar tyvar then 
387              return body_ty   -- Co-vars don't appear in body_ty!
388           else 
389              return (substTyWith [tyvar] [arg_ty] body_ty) }
390   | otherwise
391   = failWithL (mkTyAppMsg fun_ty arg_ty)
392    
393 -----------------
394 lintValApp :: CoreExpr -> OutType -> OutType -> LintM OutType
395 lintValApp arg fun_ty arg_ty
396   | Just (arg,res) <- splitFunTy_maybe fun_ty
397   = do { checkTys arg arg_ty err1
398        ; return res }
399   | otherwise
400   = failWithL err2
401   where
402     err1 = mkAppMsg       fun_ty arg_ty arg
403     err2 = mkNonFunAppMsg fun_ty arg_ty arg
404 \end{code}
405
406 \begin{code}
407 checkKinds :: OutVar -> OutType -> LintM ()
408 -- Both args have had substitution applied
409 checkKinds tyvar arg_ty
410         -- Arg type might be boxed for a function with an uncommitted
411         -- tyvar; notably this is used so that we can give
412         --      error :: forall a:*. String -> a
413         -- and then apply it to both boxed and unboxed types.
414   | isCoVar tyvar = do { (s2,t2) <- lintCoercion arg_ty
415                        ; unless (s1 `coreEqType` s2 && t1 `coreEqType` t2)
416                                 (addErrL (mkCoAppErrMsg tyvar arg_ty)) }
417   | otherwise     = do { arg_kind <- lintType arg_ty
418                        ; unless (arg_kind `isSubKind` tyvar_kind)
419                                 (addErrL (mkKindErrMsg tyvar arg_ty)) }
420   where
421     tyvar_kind = tyVarKind tyvar
422     (s1,t1)    = coVarKind tyvar
423
424 checkDeadIdOcc :: Id -> LintM ()
425 -- Occurrences of an Id should never be dead....
426 -- except when we are checking a case pattern
427 checkDeadIdOcc id
428   | isDeadOcc (idOccInfo id)
429   = do { in_case <- inCasePat
430        ; checkL in_case
431                 (ptext (sLit "Occurrence of a dead Id") <+> ppr id) }
432   | otherwise
433   = return ()
434 \end{code}
435
436
437 %************************************************************************
438 %*                                                                      *
439 \subsection[lintCoreAlts]{lintCoreAlts}
440 %*                                                                      *
441 %************************************************************************
442
443 \begin{code}
444 checkCaseAlts :: CoreExpr -> OutType -> [CoreAlt] -> LintM ()
445 -- a) Check that the alts are non-empty
446 -- b1) Check that the DEFAULT comes first, if it exists
447 -- b2) Check that the others are in increasing order
448 -- c) Check that there's a default for infinite types
449 -- NB: Algebraic cases are not necessarily exhaustive, because
450 --     the simplifer correctly eliminates case that can't 
451 --     possibly match.
452
453 checkCaseAlts e _ []
454   = addErrL (mkNullAltsMsg e)
455
456 checkCaseAlts e ty alts = 
457   do { checkL (all non_deflt con_alts) (mkNonDefltMsg e)
458      ; checkL (increasing_tag con_alts) (mkNonIncreasingAltsMsg e)
459      ; checkL (isJust maybe_deflt || not is_infinite_ty)
460            (nonExhaustiveAltsMsg e) }
461   where
462     (con_alts, maybe_deflt) = findDefault alts
463
464         -- Check that successive alternatives have increasing tags 
465     increasing_tag (alt1 : rest@( alt2 : _)) = alt1 `ltAlt` alt2 && increasing_tag rest
466     increasing_tag _                         = True
467
468     non_deflt (DEFAULT, _, _) = False
469     non_deflt _               = True
470
471     is_infinite_ty = case splitTyConApp_maybe ty of
472                         Nothing         -> False
473                         Just (tycon, _) -> isPrimTyCon tycon
474 \end{code}
475
476 \begin{code}
477 checkAltExpr :: CoreExpr -> OutType -> LintM ()
478 checkAltExpr expr ann_ty
479   = do { actual_ty <- lintCoreExpr expr 
480        ; checkTys actual_ty ann_ty (mkCaseAltMsg expr actual_ty ann_ty) }
481
482 lintCoreAlt :: OutType          -- Type of scrutinee
483             -> OutType          -- Type of the alternative
484             -> CoreAlt
485             -> LintM ()
486
487 lintCoreAlt _ alt_ty (DEFAULT, args, rhs) =
488   do { checkL (null args) (mkDefaultArgsMsg args)
489      ; checkAltExpr rhs alt_ty }
490
491 lintCoreAlt scrut_ty alt_ty (LitAlt lit, args, rhs) = 
492   do { checkL (null args) (mkDefaultArgsMsg args)
493      ; checkTys lit_ty scrut_ty (mkBadPatMsg lit_ty scrut_ty)   
494      ; checkAltExpr rhs alt_ty } 
495   where
496     lit_ty = literalType lit
497
498 lintCoreAlt scrut_ty alt_ty alt@(DataAlt con, args, rhs)
499   | isNewTyCon (dataConTyCon con) 
500   = addErrL (mkNewTyDataConAltMsg scrut_ty alt)
501   | Just (tycon, tycon_arg_tys) <- splitTyConApp_maybe scrut_ty
502   = addLoc (CaseAlt alt) $  do
503     {   -- First instantiate the universally quantified 
504         -- type variables of the data constructor
505         -- We've already check
506       checkL (tycon == dataConTyCon con) (mkBadConMsg tycon con)
507     ; let con_payload_ty = applyTys (dataConRepType con) tycon_arg_tys
508
509         -- And now bring the new binders into scope
510     ; lintBinders args $ \ args' -> do
511     { addLoc (CasePat alt) (lintAltBinders scrut_ty con_payload_ty args')
512     ; checkAltExpr rhs alt_ty } }
513
514   | otherwise   -- Scrut-ty is wrong shape
515   = addErrL (mkBadAltMsg scrut_ty alt)
516 \end{code}
517
518 %************************************************************************
519 %*                                                                      *
520 \subsection[lint-types]{Types}
521 %*                                                                      *
522 %************************************************************************
523
524 \begin{code}
525 -- When we lint binders, we (one at a time and in order):
526 --  1. Lint var types or kinds (possibly substituting)
527 --  2. Add the binder to the in scope set, and if its a coercion var,
528 --     we may extend the substitution to reflect its (possibly) new kind
529 lintBinders :: [Var] -> ([Var] -> LintM a) -> LintM a
530 lintBinders [] linterF = linterF []
531 lintBinders (var:vars) linterF = lintBinder var $ \var' ->
532                                  lintBinders vars $ \ vars' ->
533                                  linterF (var':vars')
534
535 lintBinder :: Var -> (Var -> LintM a) -> LintM a
536 lintBinder var linterF
537   | isId var  = lintIdBndr var linterF
538   | otherwise = lintTyBndr var linterF
539
540 lintTyBndr :: InTyVar -> (OutTyVar -> LintM a) -> LintM a
541 lintTyBndr tv thing_inside
542   = do { subst <- getTvSubst
543        ; let (subst', tv') = substTyVarBndr subst tv
544        ; lintTyBndrKind tv'
545        ; updateTvSubst subst' (thing_inside tv') }
546
547 lintIdBndr :: Id -> (Id -> LintM a) -> LintM a
548 -- Do substitution on the type of a binder and add the var with this 
549 -- new type to the in-scope set of the second argument
550 -- ToDo: lint its rules
551
552 lintIdBndr id linterF 
553   = do  { checkL (not (isUnboxedTupleType (idType id))) 
554                  (mkUnboxedTupleMsg id)
555                 -- No variable can be bound to an unboxed tuple.
556         ; lintAndScopeId id $ \id' -> linterF id' }
557
558 lintAndScopeIds :: [Var] -> ([Var] -> LintM a) -> LintM a
559 lintAndScopeIds ids linterF 
560   = go ids
561   where
562     go []       = linterF []
563     go (id:ids) = lintAndScopeId id $ \id ->
564                   lintAndScopeIds ids $ \ids ->
565                   linterF (id:ids)
566
567 lintAndScopeId :: InVar -> (OutVar -> LintM a) -> LintM a
568 lintAndScopeId id linterF 
569   = do { ty <- lintInTy (idType id)
570        ; let id' = setIdType id ty
571        ; addInScopeVar id' $ (linterF id') }
572 \end{code}
573
574
575 %************************************************************************
576 %*                                                                      *
577 \subsection[lint-monad]{The Lint monad}
578 %*                                                                      *
579 %************************************************************************
580
581 \begin{code}
582 lintInTy :: InType -> LintM OutType
583 -- Check the type, and apply the substitution to it
584 -- See Note [Linting type lets]
585 -- ToDo: check the kind structure of the type
586 lintInTy ty 
587   = addLoc (InType ty) $
588     do  { ty' <- applySubst ty
589         ; _ <- lintType ty'
590         ; return ty' }
591
592 -------------------
593 lintKind :: Kind -> LintM ()
594 -- Check well-formedness of kinds: *, *->*, etc
595 lintKind (TyConApp tc []) 
596   | getUnique tc `elem` kindKeys
597   = return ()
598 lintKind (FunTy k1 k2)
599   = lintKind k1 >> lintKind k2
600 lintKind kind 
601   = addErrL (hang (ptext (sLit "Malformed kind:")) 2 (quotes (ppr kind)))
602
603 -------------------
604 lintTyBndrKind :: OutTyVar -> LintM ()
605 lintTyBndrKind tv 
606   | isCoVar tv = lintCoVarKind tv
607   | otherwise  = lintKind (tyVarKind tv)
608
609 -------------------
610 lintCoVarKind :: OutCoVar -> LintM ()
611 -- Check the kind of a coercion binder
612 lintCoVarKind tv
613   = do { (ty1,ty2) <- lintSplitCoVar tv
614        ; k1 <- lintType ty1
615        ; k2 <- lintType ty2
616        ; unless (k1 `eqKind` k2) 
617                 (addErrL (sep [ ptext (sLit "Kind mis-match in coercion kind of:")
618                               , nest 2 (quotes (ppr tv))
619                               , ppr [k1,k2] ])) }
620
621 -------------------
622 lintSplitCoVar :: CoVar -> LintM (Type,Type)
623 lintSplitCoVar cv
624   = case coVarKind_maybe cv of
625       Just ts -> return ts
626       Nothing -> failWithL (sep [ ptext (sLit "Coercion variable with non-equality kind:")
627                                 , nest 2 (ppr cv <+> dcolon <+> ppr (tyVarKind cv))])
628
629 -------------------
630 lintCoercion, lintCoercion' :: OutType -> LintM (OutType, OutType)
631 -- Check the kind of a coercion term, returning the kind
632 lintCoercion co 
633   = addLoc (InCoercion co) $ lintCoercion' co
634
635 lintCoercion' ty@(TyVarTy tv)
636   = do { checkTyVarInScope tv
637        ; if isCoVar tv then return (coVarKind tv) 
638                        else return (ty, ty) }
639
640 lintCoercion' ty@(AppTy ty1 ty2) 
641   = do { (s1,t1) <- lintCoercion ty1
642        ; (s2,t2) <- lintCoercion ty2
643        ; check_co_app ty (typeKind s1) [s2]
644        ; return (mkAppTy s1 s2, mkAppTy t1 t2) }
645
646 lintCoercion' ty@(FunTy ty1 ty2)
647   = do { (s1,t1) <- lintCoercion ty1
648        ; (s2,t2) <- lintCoercion ty2
649        ; check_co_app ty (tyConKind funTyCon) [s1, s2]
650        ; return (FunTy s1 s2, FunTy t1 t2) }
651
652 lintCoercion' ty@(TyConApp tc tys) 
653   | Just (ar, desc) <- isCoercionTyCon_maybe tc
654   = do { unless (tys `lengthAtLeast` ar) (badCo ty)
655        ; (s,t) <- lintCoTyConApp ty desc (take ar tys)
656        ; (ss,ts) <- mapAndUnzipM lintCoercion (drop ar tys)
657        ; check_co_app ty (typeKind s) ss
658        ; return (mkAppTys s ss, mkAppTys t ts) }
659
660   | not (tyConHasKind tc)       -- Just something bizarre like SuperKindTyCon
661   = badCo ty
662
663   | otherwise
664   = do { (ss,ts) <- mapAndUnzipM lintCoercion tys
665        ; check_co_app ty (tyConKind tc) ss
666        ; return (TyConApp tc ss, TyConApp tc ts) }
667
668 lintCoercion' ty@(PredTy (ClassP cls tys))
669   = do { (ss,ts) <- mapAndUnzipM lintCoercion tys
670        ; check_co_app ty (tyConKind (classTyCon cls)) ss
671        ; return (PredTy (ClassP cls ss), PredTy (ClassP cls ts)) }
672
673 lintCoercion' (PredTy (IParam n p_ty))
674   = do { (s,t) <- lintCoercion p_ty
675        ; return (PredTy (IParam n s), PredTy (IParam n t)) }
676
677 lintCoercion' ty@(PredTy (EqPred {}))
678   = failWithL (badEq ty)
679
680 lintCoercion' (ForAllTy tv ty)
681   | isCoVar tv
682   = do { (co1, co2) <- lintSplitCoVar tv
683        ; (s1,t1)    <- lintCoercion co1
684        ; (s2,t2)    <- lintCoercion co2
685        ; (sr,tr)    <- lintCoercion ty
686        ; return (mkCoPredTy s1 s2 sr, mkCoPredTy t1 t2 tr) }
687
688   | otherwise
689   = do { lintKind (tyVarKind tv)
690        ; (s,t) <- addInScopeVar tv (lintCoercion ty)
691        ; return (ForAllTy tv s, ForAllTy tv t) }
692
693 badCo :: Coercion -> LintM a
694 badCo co = failWithL (hang (ptext (sLit "Ill-kinded coercion term:")) 2 (ppr co))
695
696 ---------------
697 lintCoTyConApp :: Coercion -> CoTyConDesc -> [Coercion] -> LintM (Type,Type)
698 -- Always called with correct number of coercion arguments
699 -- First arg is just for error message
700 lintCoTyConApp _ CoLeft  (co:_) = lintLR   fst      co 
701 lintCoTyConApp _ CoRight (co:_) = lintLR   snd      co   
702 lintCoTyConApp _ CoCsel1 (co:_) = lintCsel fstOf3   co 
703 lintCoTyConApp _ CoCsel2 (co:_) = lintCsel sndOf3   co 
704 lintCoTyConApp _ CoCselR (co:_) = lintCsel thirdOf3 co 
705
706 lintCoTyConApp _ CoSym (co:_) 
707   = do { (ty1,ty2) <- lintCoercion co
708        ; return (ty2,ty1) }
709
710 lintCoTyConApp co CoTrans (co1:co2:_) 
711   = do { (ty1a, ty1b) <- lintCoercion co1
712        ; (ty2a, ty2b) <- lintCoercion co2
713        ; checkL (ty1b `coreEqType` ty2a)
714                 (hang (ptext (sLit "Trans coercion mis-match:") <+> ppr co)
715                     2 (vcat [ppr ty1a, ppr ty1b, ppr ty2a, ppr ty2b]))
716        ; return (ty1a, ty2b) }
717
718 lintCoTyConApp _ CoInst (co:arg_ty:_) 
719   = do { co_tys <- lintCoercion co
720        ; arg_kind  <- lintType arg_ty
721        ; case decompInst_maybe co_tys of
722           Just ((tv1,tv2), (ty1,ty2)) 
723             | arg_kind `isSubKind` tyVarKind tv1
724             -> return (substTyWith [tv1] [arg_ty] ty1, 
725                        substTyWith [tv2] [arg_ty] ty2) 
726             | otherwise
727             -> failWithL (ptext (sLit "Kind mis-match in inst coercion"))
728           Nothing -> failWithL (ptext (sLit "Bad argument of inst")) }
729
730 lintCoTyConApp _ (CoAxiom { co_ax_tvs = tvs 
731                           , co_ax_lhs = lhs_ty, co_ax_rhs = rhs_ty }) cos
732   = do { (tys1, tys2) <- mapAndUnzipM lintCoercion cos
733        ; sequence_ (zipWith checkKinds tvs tys1)
734        ; return (substTyWith tvs tys1 lhs_ty,
735                  substTyWith tvs tys2 rhs_ty) }
736
737 lintCoTyConApp _ CoUnsafe (ty1:ty2:_) 
738   = do { _ <- lintType ty1
739        ; _ <- lintType ty2      -- Ignore kinds; it's unsafe!
740        ; return (ty1,ty2) } 
741
742 lintCoTyConApp _ _ _ = panic "lintCoTyConApp"  -- Called with wrong number of coercion args
743
744 ----------
745 lintLR :: (forall a. (a,a)->a) -> Coercion -> LintM (Type,Type)
746 lintLR sel co
747   = do { (ty1,ty2) <- lintCoercion co
748        ; case decompLR_maybe (ty1,ty2) of
749            Just res -> return (sel res)
750            Nothing  -> failWithL (ptext (sLit "Bad argument of left/right")) }
751
752 ----------
753 lintCsel :: (forall a. (a,a,a)->a) -> Coercion -> LintM (Type,Type)
754 lintCsel sel co
755   = do { (ty1,ty2) <- lintCoercion co
756        ; case decompCsel_maybe (ty1,ty2) of
757            Just res -> return (sel res)
758            Nothing  -> failWithL (ptext (sLit "Bad argument of csel")) }
759
760 -------------------
761 lintType :: OutType -> LintM Kind
762 lintType (TyVarTy tv)
763   = do { checkTyVarInScope tv
764        ; return (tyVarKind tv) }
765
766 lintType ty@(AppTy t1 t2) 
767   = do { k1 <- lintType t1
768        ; lint_ty_app ty k1 [t2] }
769
770 lintType ty@(FunTy t1 t2)
771   = lint_ty_app ty (tyConKind funTyCon) [t1,t2]
772
773 lintType ty@(TyConApp tc tys)
774   | tyConHasKind tc
775   = lint_ty_app ty (tyConKind tc) tys
776   | otherwise
777   = failWithL (hang (ptext (sLit "Malformed type:")) 2 (ppr ty))
778
779 lintType (ForAllTy tv ty)
780   = do { lintTyBndrKind tv
781        ; addInScopeVar tv (lintType ty) }
782
783 lintType ty@(PredTy (ClassP cls tys))
784   = lint_ty_app ty (tyConKind (classTyCon cls)) tys
785
786 lintType (PredTy (IParam _ p_ty))
787   = lintType p_ty
788
789 lintType ty@(PredTy (EqPred {}))
790   = failWithL (badEq ty)
791
792 ----------------
793 lint_ty_app :: Type -> Kind -> [OutType] -> LintM Kind
794 lint_ty_app ty k tys 
795   = do { ks <- mapM lintType tys
796        ; lint_kind_app (ptext (sLit "type") <+> quotes (ppr ty)) k ks }
797                       
798 ----------------
799 check_co_app :: Coercion -> Kind -> [OutType] -> LintM ()
800 check_co_app ty k tys 
801   = do { _ <- lint_kind_app (ptext (sLit "coercion") <+> quotes (ppr ty))  
802                             k (map typeKind tys)
803        ; return () }
804                       
805 ----------------
806 lint_kind_app :: SDoc -> Kind -> [Kind] -> LintM Kind
807 lint_kind_app doc kfn ks = go kfn ks
808   where
809     fail_msg = vcat [hang (ptext (sLit "Kind application error in")) 2 doc,
810                      nest 2 (ptext (sLit "Function kind =") <+> ppr kfn),
811                      nest 2 (ptext (sLit "Arg kinds =") <+> ppr ks)]
812
813     go kfn []     = return kfn
814     go kfn (k:ks) = case splitKindFunTy_maybe kfn of
815                       Nothing         -> failWithL fail_msg
816                       Just (kfa, kfb) -> do { unless (k `isSubKind` kfa)
817                                                      (addErrL fail_msg)
818                                             ; go kfb ks } 
819 --------------
820 badEq :: Type -> SDoc
821 badEq ty = hang (ptext (sLit "Unexpected equality predicate:"))
822               1 (quotes (ppr ty))
823 \end{code}
824     
825 %************************************************************************
826 %*                                                                      *
827 \subsection[lint-monad]{The Lint monad}
828 %*                                                                      *
829 %************************************************************************
830
831 \begin{code}
832 newtype LintM a = 
833    LintM { unLintM :: 
834             [LintLocInfo] ->         -- Locations
835             TvSubst ->               -- Current type substitution; we also use this
836                                      -- to keep track of all the variables in scope,
837                                      -- both Ids and TyVars
838             WarnsAndErrs ->           -- Error and warning messages so far
839             (Maybe a, WarnsAndErrs) } -- Result and messages (if any)
840
841 type WarnsAndErrs = (Bag Message, Bag Message)
842
843 {-      Note [Type substitution]
844         ~~~~~~~~~~~~~~~~~~~~~~~~
845 Why do we need a type substitution?  Consider
846         /\(a:*). \(x:a). /\(a:*). id a x
847 This is ill typed, because (renaming variables) it is really
848         /\(a:*). \(x:a). /\(b:*). id b x
849 Hence, when checking an application, we can't naively compare x's type
850 (at its binding site) with its expected type (at a use site).  So we
851 rename type binders as we go, maintaining a substitution.
852
853 The same substitution also supports let-type, current expressed as
854         (/\(a:*). body) ty
855 Here we substitute 'ty' for 'a' in 'body', on the fly.
856 -}
857
858 instance Monad LintM where
859   return x = LintM (\ _   _     errs -> (Just x, errs))
860   fail err = failWithL (text err)
861   m >>= k  = LintM (\ loc subst errs -> 
862                        let (res, errs') = unLintM m loc subst errs in
863                          case res of
864                            Just r -> unLintM (k r) loc subst errs'
865                            Nothing -> (Nothing, errs'))
866
867 data LintLocInfo
868   = RhsOf Id            -- The variable bound
869   | LambdaBodyOf Id     -- The lambda-binder
870   | BodyOfLetRec [Id]   -- One of the binders
871   | CaseAlt CoreAlt     -- Case alternative
872   | CasePat CoreAlt     -- The *pattern* of the case alternative
873   | AnExpr CoreExpr     -- Some expression
874   | ImportedUnfolding SrcLoc -- Some imported unfolding (ToDo: say which)
875   | TopLevelBindings
876   | InType Type         -- Inside a type
877   | InCoercion Coercion -- Inside a type
878 \end{code}
879
880                  
881 \begin{code}
882 initL :: LintM a -> WarnsAndErrs    -- Errors and warnings
883 initL m
884   = case unLintM m [] emptyTvSubst (emptyBag, emptyBag) of
885       (_, errs) -> errs
886 \end{code}
887
888 \begin{code}
889 checkL :: Bool -> Message -> LintM ()
890 checkL True  _   = return ()
891 checkL False msg = failWithL msg
892
893 failWithL :: Message -> LintM a
894 failWithL msg = LintM $ \ loc subst (warns,errs) ->
895                 (Nothing, (warns, addMsg subst errs msg loc))
896
897 addErrL :: Message -> LintM ()
898 addErrL msg = LintM $ \ loc subst (warns,errs) -> 
899               (Just (), (warns, addMsg subst errs msg loc))
900
901 addWarnL :: Message -> LintM ()
902 addWarnL msg = LintM $ \ loc subst (warns,errs) -> 
903               (Just (), (addMsg subst warns msg loc, errs))
904
905 addMsg :: TvSubst ->  Bag Message -> Message -> [LintLocInfo] -> Bag Message
906 addMsg subst msgs msg locs
907   = ASSERT( notNull locs )
908     msgs `snocBag` mk_msg msg
909   where
910    (loc, cxt1) = dumpLoc (head locs)
911    cxts        = [snd (dumpLoc loc) | loc <- locs]   
912    context     | opt_PprStyle_Debug = vcat (reverse cxts) $$ cxt1 $$
913                                       ptext (sLit "Substitution:") <+> ppr subst
914                | otherwise          = cxt1
915  
916    mk_msg msg = mkLocMessage (mkSrcSpan loc loc) (context $$ msg)
917
918 addLoc :: LintLocInfo -> LintM a -> LintM a
919 addLoc extra_loc m =
920   LintM (\ loc subst errs -> unLintM m (extra_loc:loc) subst errs)
921
922 inCasePat :: LintM Bool         -- A slight hack; see the unique call site
923 inCasePat = LintM $ \ loc _ errs -> (Just (is_case_pat loc), errs)
924   where
925     is_case_pat (CasePat {} : _) = True
926     is_case_pat _other           = False
927
928 addInScopeVars :: [Var] -> LintM a -> LintM a
929 addInScopeVars vars m
930   = LintM (\ loc subst errs -> unLintM m loc (extendTvInScopeList subst vars) errs)
931
932 addInScopeVar :: Var -> LintM a -> LintM a
933 addInScopeVar var m
934   = LintM (\ loc subst errs -> unLintM m loc (extendTvInScope subst var) errs)
935
936 updateTvSubst :: TvSubst -> LintM a -> LintM a
937 updateTvSubst subst' m = 
938   LintM (\ loc _ errs -> unLintM m loc subst' errs)
939
940 getTvSubst :: LintM TvSubst
941 getTvSubst = LintM (\ _ subst errs -> (Just subst, errs))
942
943 applySubst :: Type -> LintM Type
944 applySubst ty = do { subst <- getTvSubst; return (substTy subst ty) }
945
946 extendSubstL :: TyVar -> Type -> LintM a -> LintM a
947 extendSubstL tv ty m
948   = LintM (\ loc subst errs -> unLintM m loc (extendTvSubst subst tv ty) errs)
949 \end{code}
950
951 \begin{code}
952 lookupIdInScope :: Id -> LintM Id
953 lookupIdInScope id 
954   | not (mustHaveLocalBinding id)
955   = return id   -- An imported Id
956   | otherwise   
957   = do  { subst <- getTvSubst
958         ; case lookupInScope (getTvInScope subst) id of
959                 Just v  -> return v
960                 Nothing -> do { addErrL out_of_scope
961                               ; return id } }
962   where
963     out_of_scope = ppr id <+> ptext (sLit "is out of scope")
964
965
966 oneTupleDataConId :: Id -- Should not happen
967 oneTupleDataConId = dataConWorkId (tupleCon Boxed 1)
968
969 checkBndrIdInScope :: Var -> Var -> LintM ()
970 checkBndrIdInScope binder id 
971   = checkInScope msg id
972     where
973      msg = ptext (sLit "is out of scope inside info for") <+> 
974            ppr binder
975
976 checkTyVarInScope :: TyVar -> LintM ()
977 checkTyVarInScope tv = checkInScope (ptext (sLit "is out of scope")) tv
978
979 checkInScope :: SDoc -> Var -> LintM ()
980 checkInScope loc_msg var =
981  do { subst <- getTvSubst
982     ; checkL (not (mustHaveLocalBinding var) || (var `isInScope` subst))
983              (hsep [ppr var, loc_msg]) }
984
985 checkTys :: OutType -> OutType -> Message -> LintM ()
986 -- check ty2 is subtype of ty1 (ie, has same structure but usage
987 -- annotations need only be consistent, not equal)
988 -- Assumes ty1,ty2 are have alrady had the substitution applied
989 checkTys ty1 ty2 msg = checkL (ty1 `coreEqType` ty2) msg
990 \end{code}
991
992 %************************************************************************
993 %*                                                                      *
994 \subsection{Error messages}
995 %*                                                                      *
996 %************************************************************************
997
998 \begin{code}
999 dumpLoc :: LintLocInfo -> (SrcLoc, SDoc)
1000
1001 dumpLoc (RhsOf v)
1002   = (getSrcLoc v, brackets (ptext (sLit "RHS of") <+> pp_binders [v]))
1003
1004 dumpLoc (LambdaBodyOf b)
1005   = (getSrcLoc b, brackets (ptext (sLit "in body of lambda with binder") <+> pp_binder b))
1006
1007 dumpLoc (BodyOfLetRec [])
1008   = (noSrcLoc, brackets (ptext (sLit "In body of a letrec with no binders")))
1009
1010 dumpLoc (BodyOfLetRec bs@(_:_))
1011   = ( getSrcLoc (head bs), brackets (ptext (sLit "in body of letrec with binders") <+> pp_binders bs))
1012
1013 dumpLoc (AnExpr e)
1014   = (noSrcLoc, text "In the expression:" <+> ppr e)
1015
1016 dumpLoc (CaseAlt (con, args, _))
1017   = (noSrcLoc, text "In a case alternative:" <+> parens (ppr con <+> pp_binders args))
1018
1019 dumpLoc (CasePat (con, args, _))
1020   = (noSrcLoc, text "In the pattern of a case alternative:" <+> parens (ppr con <+> pp_binders args))
1021
1022 dumpLoc (ImportedUnfolding locn)
1023   = (locn, brackets (ptext (sLit "in an imported unfolding")))
1024 dumpLoc TopLevelBindings
1025   = (noSrcLoc, empty)
1026 dumpLoc (InType ty)
1027   = (noSrcLoc, text "In the type" <+> quotes (ppr ty))
1028 dumpLoc (InCoercion ty)
1029   = (noSrcLoc, text "In the coercion" <+> quotes (ppr ty))
1030
1031 pp_binders :: [Var] -> SDoc
1032 pp_binders bs = sep (punctuate comma (map pp_binder bs))
1033
1034 pp_binder :: Var -> SDoc
1035 pp_binder b | isId b    = hsep [ppr b, dcolon, ppr (idType b)]
1036             | otherwise = hsep [ppr b, dcolon, ppr (tyVarKind b)]
1037 \end{code}
1038
1039 \begin{code}
1040 ------------------------------------------------------
1041 --      Messages for case expressions
1042
1043 mkNullAltsMsg :: CoreExpr -> Message
1044 mkNullAltsMsg e 
1045   = hang (text "Case expression with no alternatives:")
1046          4 (ppr e)
1047
1048 mkDefaultArgsMsg :: [Var] -> Message
1049 mkDefaultArgsMsg args 
1050   = hang (text "DEFAULT case with binders")
1051          4 (ppr args)
1052
1053 mkCaseAltMsg :: CoreExpr -> Type -> Type -> Message
1054 mkCaseAltMsg e ty1 ty2
1055   = hang (text "Type of case alternatives not the same as the annotation on case:")
1056          4 (vcat [ppr ty1, ppr ty2, ppr e])
1057
1058 mkScrutMsg :: Id -> Type -> Type -> TvSubst -> Message
1059 mkScrutMsg var var_ty scrut_ty subst
1060   = vcat [text "Result binder in case doesn't match scrutinee:" <+> ppr var,
1061           text "Result binder type:" <+> ppr var_ty,--(idType var),
1062           text "Scrutinee type:" <+> ppr scrut_ty,
1063      hsep [ptext (sLit "Current TV subst"), ppr subst]]
1064
1065 mkNonDefltMsg, mkNonIncreasingAltsMsg :: CoreExpr -> Message
1066 mkNonDefltMsg e
1067   = hang (text "Case expression with DEFAULT not at the beginnning") 4 (ppr e)
1068 mkNonIncreasingAltsMsg e
1069   = hang (text "Case expression with badly-ordered alternatives") 4 (ppr e)
1070
1071 nonExhaustiveAltsMsg :: CoreExpr -> Message
1072 nonExhaustiveAltsMsg e
1073   = hang (text "Case expression with non-exhaustive alternatives") 4 (ppr e)
1074
1075 mkBadConMsg :: TyCon -> DataCon -> Message
1076 mkBadConMsg tycon datacon
1077   = vcat [
1078         text "In a case alternative, data constructor isn't in scrutinee type:",
1079         text "Scrutinee type constructor:" <+> ppr tycon,
1080         text "Data con:" <+> ppr datacon
1081     ]
1082
1083 mkBadPatMsg :: Type -> Type -> Message
1084 mkBadPatMsg con_result_ty scrut_ty
1085   = vcat [
1086         text "In a case alternative, pattern result type doesn't match scrutinee type:",
1087         text "Pattern result type:" <+> ppr con_result_ty,
1088         text "Scrutinee type:" <+> ppr scrut_ty
1089     ]
1090
1091 mkBadAltMsg :: Type -> CoreAlt -> Message
1092 mkBadAltMsg scrut_ty alt
1093   = vcat [ text "Data alternative when scrutinee is not a tycon application",
1094            text "Scrutinee type:" <+> ppr scrut_ty,
1095            text "Alternative:" <+> pprCoreAlt alt ]
1096
1097 mkNewTyDataConAltMsg :: Type -> CoreAlt -> Message
1098 mkNewTyDataConAltMsg scrut_ty alt
1099   = vcat [ text "Data alternative for newtype datacon",
1100            text "Scrutinee type:" <+> ppr scrut_ty,
1101            text "Alternative:" <+> pprCoreAlt alt ]
1102
1103
1104 ------------------------------------------------------
1105 --      Other error messages
1106
1107 mkAppMsg :: Type -> Type -> CoreExpr -> Message
1108 mkAppMsg fun_ty arg_ty arg
1109   = vcat [ptext (sLit "Argument value doesn't match argument type:"),
1110               hang (ptext (sLit "Fun type:")) 4 (ppr fun_ty),
1111               hang (ptext (sLit "Arg type:")) 4 (ppr arg_ty),
1112               hang (ptext (sLit "Arg:")) 4 (ppr arg)]
1113
1114 mkNonFunAppMsg :: Type -> Type -> CoreExpr -> Message
1115 mkNonFunAppMsg fun_ty arg_ty arg
1116   = vcat [ptext (sLit "Non-function type in function position"),
1117               hang (ptext (sLit "Fun type:")) 4 (ppr fun_ty),
1118               hang (ptext (sLit "Arg type:")) 4 (ppr arg_ty),
1119               hang (ptext (sLit "Arg:")) 4 (ppr arg)]
1120
1121 mkTyVarLetErr :: TyVar -> Type -> Message
1122 mkTyVarLetErr tyvar ty
1123   = vcat [ptext (sLit "Bad `let' binding for type or coercion variable:"),
1124           hang (ptext (sLit "Type/coercion variable:"))
1125                  4 (ppr tyvar <+> dcolon <+> ppr (tyVarKind tyvar)),
1126           hang (ptext (sLit "Arg type/coercion:"))   
1127                  4 (ppr ty)]
1128
1129 mkKindErrMsg :: TyVar -> Type -> Message
1130 mkKindErrMsg tyvar arg_ty
1131   = vcat [ptext (sLit "Kinds don't match in type application:"),
1132           hang (ptext (sLit "Type variable:"))
1133                  4 (ppr tyvar <+> dcolon <+> ppr (tyVarKind tyvar)),
1134           hang (ptext (sLit "Arg type:"))   
1135                  4 (ppr arg_ty <+> dcolon <+> ppr (typeKind arg_ty))]
1136
1137 mkCoAppErrMsg :: TyVar -> Type -> Message
1138 mkCoAppErrMsg tyvar arg_ty
1139   = vcat [ptext (sLit "Kinds don't match in coercion application:"),
1140           hang (ptext (sLit "Coercion variable:"))
1141                  4 (ppr tyvar <+> dcolon <+> ppr (tyVarKind tyvar)),
1142           hang (ptext (sLit "Arg coercion:"))   
1143                  4 (ppr arg_ty <+> dcolon <+> pprEqPred (coercionKind arg_ty))]
1144
1145 mkTyAppMsg :: Type -> Type -> Message
1146 mkTyAppMsg ty arg_ty
1147   = vcat [text "Illegal type application:",
1148               hang (ptext (sLit "Exp type:"))
1149                  4 (ppr ty <+> dcolon <+> ppr (typeKind ty)),
1150               hang (ptext (sLit "Arg type:"))   
1151                  4 (ppr arg_ty <+> dcolon <+> ppr (typeKind arg_ty))]
1152
1153 mkRhsMsg :: Id -> Type -> Message
1154 mkRhsMsg binder ty
1155   = vcat
1156     [hsep [ptext (sLit "The type of this binder doesn't match the type of its RHS:"),
1157             ppr binder],
1158      hsep [ptext (sLit "Binder's type:"), ppr (idType binder)],
1159      hsep [ptext (sLit "Rhs type:"), ppr ty]]
1160
1161 mkRhsPrimMsg :: Id -> CoreExpr -> Message
1162 mkRhsPrimMsg binder _rhs
1163   = vcat [hsep [ptext (sLit "The type of this binder is primitive:"),
1164                      ppr binder],
1165               hsep [ptext (sLit "Binder's type:"), ppr (idType binder)]
1166              ]
1167
1168 mkStrictMsg :: Id -> Message
1169 mkStrictMsg binder
1170   = vcat [hsep [ptext (sLit "Recursive or top-level binder has strict demand info:"),
1171                      ppr binder],
1172               hsep [ptext (sLit "Binder's demand info:"), ppr (idDemandInfo binder)]
1173              ]
1174
1175 mkArityMsg :: Id -> Message
1176 mkArityMsg binder
1177   = vcat [hsep [ptext (sLit "Demand type has "),
1178                      ppr (dmdTypeDepth dmd_ty),
1179                      ptext (sLit " arguments, rhs has "),
1180                      ppr (idArity binder),
1181                      ptext (sLit "arguments, "),
1182                      ppr binder],
1183               hsep [ptext (sLit "Binder's strictness signature:"), ppr dmd_ty]
1184
1185          ]
1186            where (StrictSig dmd_ty) = idStrictness binder
1187
1188 mkUnboxedTupleMsg :: Id -> Message
1189 mkUnboxedTupleMsg binder
1190   = vcat [hsep [ptext (sLit "A variable has unboxed tuple type:"), ppr binder],
1191           hsep [ptext (sLit "Binder's type:"), ppr (idType binder)]]
1192
1193 mkCastErr :: Type -> Type -> Message
1194 mkCastErr from_ty expr_ty
1195   = vcat [ptext (sLit "From-type of Cast differs from type of enclosed expression"),
1196           ptext (sLit "From-type:") <+> ppr from_ty,
1197           ptext (sLit "Type of enclosed expr:") <+> ppr expr_ty
1198     ]
1199
1200 dupVars :: [[Var]] -> Message
1201 dupVars vars
1202   = hang (ptext (sLit "Duplicate variables brought into scope"))
1203        2 (ppr vars)
1204
1205 dupExtVars :: [[Name]] -> Message
1206 dupExtVars vars
1207   = hang (ptext (sLit "Duplicate top-level variables with the same qualified name"))
1208        2 (ppr vars)
1209 \end{code}