A radical overhaul of the coercion infrastucture
authorsimonpj@microsoft.com <unknown>
Thu, 12 Nov 2009 09:29:06 +0000 (09:29 +0000)
committersimonpj@microsoft.com <unknown>
Thu, 12 Nov 2009 09:29:06 +0000 (09:29 +0000)
* Core Lint now does full checking of kinds and coercion terms
  which picks up kind errors in coercions that were previously
  simply not checked for

* Coercion.lhs now provides optCoercion which optimises coercion
  terms.  It implements all of Dimitrios's rules

* The constructors for coercion terms now make no attempt to be
  "smart"; instead we rely solely on the coercion optimiser

* CoercionTyCons in TyCon.lhs always had a "custom" kinding rule
  (the coKindFun field of CoercionTyCon) but its type was not
  clever enough to do both
     (a) *figure out the result kind*, assuming the whole thing
         is well-kinded in the first place
     (b) *check* the kinds of everything, failing gracefully if
         they aren't right.
  We need (b) for the new CoreLint stuff. The field now has type
        CoTyConKindChecker
  which does the job nicely.

compiler/coreSyn/CoreArity.lhs
compiler/coreSyn/CoreLint.lhs
compiler/prelude/PrelNames.lhs
compiler/types/Coercion.lhs
compiler/types/TyCon.lhs
compiler/types/Type.lhs

index cacf423..673d619 100644 (file)
@@ -631,6 +631,6 @@ freshEtaId n subst ty
         ty'     = substTy subst ty
        eta_id' = uniqAway (getTvInScope subst) $
                  mkSysLocal (fsLit "eta") (mkBuiltinUnique n) ty'
-       subst'  = extendTvInScope subst [eta_id']                 
+       subst'  = extendTvInScope subst eta_id'           
 \end{code}
 
index 2747625..70bf08b 100644 (file)
@@ -28,11 +28,14 @@ import PprCore
 import ErrUtils
 import SrcLoc
 import Type
+import TypeRep
 import Coercion
 import TyCon
+import Class
 import BasicTypes
 import StaticFlags
 import ListSetOps
+import PrelNames
 import DynFlags
 import Outputable
 import FastString
@@ -81,7 +84,7 @@ That is, use a type let.   See Note [Type let] in CoreSyn.
 However, when linting <body> we need to remember that a=Int, else we might
 reject a correct program.  So we carry a type substitution (in this example 
 [a -> Int]) and apply this substitution before comparing types.  The functin
-       lintTy :: Type -> LintM Type
+       lintInTy :: Type -> LintM Type
 returns a substituted type; that's the only reason it returns anything.
 
 When we encounter a binder (like x::a) we must apply the substitution
@@ -214,7 +217,13 @@ lintSingleBinding top_lvl_flag rec_flag (binder,rhs)
 
 \begin{code}
 type InType  = Type    -- Substitution not yet applied
-type OutType = Type    -- Substitution has been applied to this
+type InVar   = Var
+type InTyVar = TyVar
+
+type OutType  = Type   -- Substitution has been applied to this
+type OutVar   = Var
+type OutTyVar = TyVar
+type OutCoVar = CoVar
 
 lintCoreExpr :: CoreExpr -> LintM OutType
 -- The returned type has the substitution from the monad 
@@ -235,17 +244,10 @@ lintCoreExpr (Var var)
 lintCoreExpr (Lit lit)
   = return (literalType lit)
 
---lintCoreExpr (Note (Coerce to_ty from_ty) expr)
---  = do       { expr_ty <- lintCoreExpr expr
---     ; to_ty <- lintTy to_ty
---     ; from_ty <- lintTy from_ty     
---     ; checkTys from_ty expr_ty (mkCoerceErr from_ty expr_ty)
---     ; return to_ty }
-
 lintCoreExpr (Cast expr co)
   = do { expr_ty <- lintCoreExpr expr
-       ; co' <- lintTy co
-       ; let (from_ty, to_ty) = coercionKind co'
+       ; co' <- applySubst co
+       ; (from_ty, to_ty) <- lintCoercion co'
        ; checkTys from_ty expr_ty (mkCastErr from_ty expr_ty)
        ; return to_ty }
 
@@ -255,16 +257,14 @@ lintCoreExpr (Note _ expr)
 lintCoreExpr (Let (NonRec tv (Type ty)) body)
   =    -- See Note [Type let] in CoreSyn
     do { checkL (isTyVar tv) (mkKindErrMsg tv ty)      -- Not quite accurate
-       ; ty' <- lintTy ty
-        ; kind' <- lintTy (tyVarKind tv)
-        ; let tv' = setTyVarKind tv kind'
-        ; checkKinds tv' ty'              
+       ; ty' <- lintInTy ty
+        ; lintTyBndr tv              $ \ tv' -> 
+          addLoc (BodyOfLetRec [tv]) $ 
+          extendSubstL tv' ty'       $ do
+        { checkKinds tv' ty'              
                -- Now extend the substitution so we 
                -- take advantage of it in the body
-        ; addLoc (BodyOfLetRec [tv]) $
-         addInScopeVars [tv'] $
-          extendSubstL tv' ty' $
-         lintCoreExpr body }
+        ; lintCoreExpr body } }
 
 lintCoreExpr (Let (NonRec bndr rhs) body)
   = do { lintSingleBinding NotTopLevel NonRecursive (bndr,rhs)
@@ -297,8 +297,8 @@ lintCoreExpr (Lam var expr)
 lintCoreExpr e@(Case scrut var alt_ty alts) =
        -- Check the scrutinee
   do { scrut_ty <- lintCoreExpr scrut
-     ; alt_ty   <- lintTy alt_ty  
-     ; var_ty   <- lintTy (idType var) 
+     ; alt_ty   <- lintInTy alt_ty  
+     ; var_ty   <- lintInTy (idType var)       
 
      ; let mb_tc_app = splitTyConApp_maybe (idType var)
      ; case mb_tc_app of 
@@ -330,7 +330,7 @@ lintCoreExpr e@(Case scrut var alt_ty alts) =
     pass_var f = f var
 
 lintCoreExpr (Type ty)
-  = do { ty' <- lintTy ty
+  = do { ty' <- lintInTy ty
        ; return (typeKind ty') }
 \end{code}
 
@@ -355,49 +355,46 @@ lintCoreArgs ty (a : args) =
   do { res <- lintCoreArg ty a
      ; lintCoreArgs res args }
 
-lintCoreArg fun_ty (Type arg_ty) =
-  do { arg_ty <- lintTy arg_ty 
-     ; lintTyApp fun_ty arg_ty }
+lintCoreArg fun_ty (Type arg_ty)
+  | Just (tyvar,body) <- splitForAllTy_maybe fun_ty
+  = do { arg_ty' <- applySubst arg_ty
+        ; checkKinds tyvar arg_ty'
+       ; if isCoVar tyvar then 
+             return body   -- Co-vars don't appear in body!
+          else 
+             return (substTyWith [tyvar] [arg_ty'] body) }
+  | otherwise
+  = failWithL (mkTyAppMsg fun_ty arg_ty)
 
-lintCoreArg fun_ty arg = 
+lintCoreArg fun_ty arg
        -- Make sure function type matches argument
-  do { arg_ty <- lintCoreExpr arg
-     ; let err1 =  mkAppMsg fun_ty arg_ty arg
-           err2 = mkNonFunAppMsg fun_ty arg_ty arg
-     ; case splitFunTy_maybe fun_ty of
-        Just (arg,res) -> 
-          do { checkTys arg arg_ty err1
-             ; return res }
-        _ -> failWithL err2 }
+ = do { arg_ty <- lintCoreExpr arg
+      ; let err1 = mkAppMsg fun_ty arg_ty arg
+            err2 = mkNonFunAppMsg fun_ty arg_ty arg
+      ; case splitFunTy_maybe fun_ty of
+          Just (arg,res) -> 
+            do { checkTys arg arg_ty err1
+               ; return res }
+          _ -> failWithL err2 }
 \end{code}
 
 \begin{code}
+checkKinds :: Var -> OutType -> LintM ()
 -- Both args have had substitution applied
-lintTyApp :: OutType -> OutType -> LintM OutType
-lintTyApp ty arg_ty 
-  = case splitForAllTy_maybe ty of
-      Nothing -> failWithL (mkTyAppMsg ty arg_ty)
-
-      Just (tyvar,body)
-        -> do  { checkL (isTyVar tyvar) (mkTyAppMsg ty arg_ty)
-               ; checkKinds tyvar arg_ty
-               ; return (substTyWith [tyvar] [arg_ty] body) }
-
-checkKinds :: Var -> Type -> LintM ()
 checkKinds tyvar arg_ty
        -- Arg type might be boxed for a function with an uncommitted
        -- tyvar; notably this is used so that we can give
        --      error :: forall a:*. String -> a
        -- and then apply it to both boxed and unboxed types.
-  | isCoVar tyvar = unless (s1 `coreEqType` s2 && t1 `coreEqType` t2)
-                           (addErrL (mkCoAppErrMsg tyvar arg_ty))
-  | otherwise     = unless (arg_kind `isSubKind` tyvar_kind)
-                           (addErrL (mkKindErrMsg tyvar arg_ty))
+  | isCoVar tyvar = do { (s2,t2) <- lintCoercion arg_ty
+                       ; unless (s1 `coreEqType` s2 && t1 `coreEqType` t2)
+                                (addErrL (mkCoAppErrMsg tyvar arg_ty)) }
+  | otherwise     = do { arg_kind <- lintType arg_ty
+                       ; unless (arg_kind `isSubKind` tyvar_kind)
+                                (addErrL (mkKindErrMsg tyvar arg_ty)) }
   where
     tyvar_kind = tyVarKind tyvar
-    arg_kind   = typeKind arg_ty
-    (s1,t1) = coVarKind tyvar
-    (s2,t2) = coercionKind arg_ty
+    (s1,t1)    = coVarKind tyvar
 
 checkDeadIdOcc :: Id -> LintM ()
 -- Occurrences of an Id should never be dead....
@@ -522,51 +519,227 @@ lintBinders (var:vars) linterF = lintBinder var $ \var' ->
 
 lintBinder :: Var -> (Var -> LintM a) -> LintM a
 lintBinder var linterF
-  | isTyVar var = lint_ty_bndr
-  | otherwise   = lintIdBndr var linterF
-  where
-    lint_ty_bndr = do { _ <- lintTy (tyVarKind var)
-                     ; subst <- getTvSubst
-                     ; let (subst', tv') = substTyVarBndr subst var
-                     ; updateTvSubst subst' (linterF tv') }
+  | isId var  = lintIdBndr var linterF
+  | otherwise = lintTyBndr var linterF
+
+lintTyBndr :: InTyVar -> (OutTyVar -> LintM a) -> LintM a
+lintTyBndr tv thing_inside
+  = do { subst <- getTvSubst
+       ; let (subst', tv') = substTyVarBndr subst tv
+       ; lintTyBndrKind tv'
+       ; updateTvSubst subst' (thing_inside tv') }
 
-lintIdBndr :: Var -> (Var -> LintM a) -> LintM a
+lintIdBndr :: Id -> (Id -> LintM a) -> LintM a
 -- Do substitution on the type of a binder and add the var with this 
 -- new type to the in-scope set of the second argument
 -- ToDo: lint its rules
+
 lintIdBndr id linterF 
   = do         { checkL (not (isUnboxedTupleType (idType id))) 
                 (mkUnboxedTupleMsg id)
                -- No variable can be bound to an unboxed tuple.
-        ; lintAndScopeId id $ \id' -> linterF id'
-        }
+        ; lintAndScopeId id $ \id' -> linterF id' }
 
 lintAndScopeIds :: [Var] -> ([Var] -> LintM a) -> LintM a
 lintAndScopeIds ids linterF 
   = go ids
   where
     go []       = linterF []
-    go (id:ids) = do { lintAndScopeId id $ \id ->
-                           lintAndScopeIds ids $ \ids ->
-                           linterF (id:ids) }
+    go (id:ids) = lintAndScopeId id $ \id ->
+                  lintAndScopeIds ids $ \ids ->
+                  linterF (id:ids)
 
-lintAndScopeId :: Var -> (Var -> LintM a) -> LintM a
+lintAndScopeId :: InVar -> (OutVar -> LintM a) -> LintM a
 lintAndScopeId id linterF 
-  = do { ty <- lintTy (idType id)
+  = do { ty <- lintInTy (idType id)
        ; let id' = setIdType id ty
-       ; addInScopeVars [id'] $ (linterF id')
-       }
+       ; addInScopeVar id' $ (linterF id') }
+\end{code}
+
+
+%************************************************************************
+%*                                                                     *
+\subsection[lint-monad]{The Lint monad}
+%*                                                                     *
+%************************************************************************
 
-lintTy :: InType -> LintM OutType
+\begin{code}
+lintInTy :: InType -> LintM OutType
 -- Check the type, and apply the substitution to it
 -- See Note [Linting type lets]
 -- ToDo: check the kind structure of the type
-lintTy ty 
-  = do { ty' <- applySubst ty
-       ; mapM_ checkTyVarInScope (varSetElems (tyVarsOfType ty'))
+lintInTy ty 
+  = addLoc (InType ty) $
+    do { ty' <- applySubst ty
+       ; _ <- lintType ty'
        ; return ty' }
-\end{code}
 
+-------------------
+lintKind :: Kind -> LintM ()
+-- Check well-formedness of kinds: *, *->*, etc
+lintKind (TyConApp tc []) 
+  | getUnique tc `elem` kindKeys
+  = return ()
+lintKind (FunTy k1 k2)
+  = lintKind k1 >> lintKind k2
+lintKind kind 
+  = addErrL (hang (ptext (sLit "Malformed kind:")) 2 (quotes (ppr kind)))
+
+-------------------
+lintTyBndrKind :: OutTyVar -> LintM ()
+lintTyBndrKind tv 
+  | isCoVar tv = lintCoVarKind tv
+  | otherwise  = lintKind (tyVarKind tv)
+
+-------------------
+lintCoVarKind :: OutCoVar -> LintM ()
+-- Check the kind of a coercion binder
+lintCoVarKind tv
+  = do { (ty1,ty2) <- lintSplitCoVar tv
+       ; k1 <- lintType ty1
+       ; k2 <- lintType ty2
+       ; unless (k1 `eqKind` k2) 
+                (addErrL (sep [ ptext (sLit "Kind mis-match in coercion kind of:")
+                              , nest 2 (quotes (ppr tv))
+                              , ppr [k1,k2] ])) }
+
+-------------------
+lintSplitCoVar :: CoVar -> LintM (Type,Type)
+lintSplitCoVar cv
+  = case coVarKind_maybe cv of
+      Just ts -> return ts
+      Nothing -> failWithL (sep [ ptext (sLit "Coercion variable with non-equality kind:")
+                                , nest 2 (ppr cv <+> dcolon <+> ppr (tyVarKind cv))])
+
+-------------------
+lintCoercion :: OutType -> LintM (OutType, OutType)
+-- Check the kind of a coercion term, returning the kind
+lintCoercion ty@(TyVarTy tv)
+  = do { checkTyVarInScope tv
+       ; if isCoVar tv then return (coVarKind tv) 
+                       else return (ty, ty) }
+
+lintCoercion ty@(AppTy ty1 ty2) 
+  = do { (s1,t1) <- lintCoercion ty1
+       ; (s2,t2) <- lintCoercion ty2
+       ; check_co_app ty (typeKind s1) [s2]
+       ; return (AppTy s1 s2, AppTy t1 t2) }
+
+lintCoercion ty@(FunTy ty1 ty2) 
+  = do { (s1,t1) <- lintCoercion ty1
+       ; (s2,t2) <- lintCoercion ty2
+       ; check_co_app ty (tyConKind funTyCon) [s1, s2]
+       ; return (FunTy s1 s2, FunTy t1 t2) }
+
+lintCoercion ty@(TyConApp tc tys) 
+  | Just (ar, rule) <- isCoercionTyCon_maybe tc
+  = do { unless (tys `lengthAtLeast` ar) (badCo ty)
+       ; (s,t)   <- rule lintType lintCoercion 
+                         True (take ar tys)
+       ; (ss,ts) <- mapAndUnzipM lintCoercion (drop ar tys)
+       ; check_co_app ty (typeKind s) ss
+       ; return (mkAppTys s ss, mkAppTys t ts) }
+
+  | not (tyConHasKind tc)      -- Just something bizarre like SuperKindTyCon
+  = badCo ty
+
+  | otherwise
+  = do { (ss,ts) <- mapAndUnzipM lintCoercion tys
+       ; check_co_app ty (tyConKind tc) ss
+       ; return (TyConApp tc ss, TyConApp tc ts) }
+
+lintCoercion ty@(PredTy (ClassP cls tys))
+  = do { (ss,ts) <- mapAndUnzipM lintCoercion tys
+       ; check_co_app ty (tyConKind (classTyCon cls)) ss
+       ; return (PredTy (ClassP cls ss), PredTy (ClassP cls ts)) }
+
+lintCoercion (PredTy (IParam n p_ty))
+  = do { (s,t) <- lintCoercion p_ty
+       ; return (PredTy (IParam n s), PredTy (IParam n t)) }
+
+lintCoercion ty@(PredTy (EqPred {}))
+  = failWithL (badEq ty)
+
+lintCoercion (ForAllTy tv ty)
+  | isCoVar tv
+  = do { (co1, co2) <- lintSplitCoVar tv
+       ; (s1,t1)    <- lintCoercion co1
+       ; (s2,t2)    <- lintCoercion co2
+       ; (sr,tr)    <- lintCoercion ty
+       ; return (mkCoPredTy s1 s2 sr, mkCoPredTy t1 t2 tr) }
+
+  | otherwise
+  = do { lintKind (tyVarKind tv)
+       ; (s,t) <- addInScopeVar tv (lintCoercion ty)
+       ; return (ForAllTy tv s, ForAllTy tv t) }
+
+badCo :: Coercion -> LintM a
+badCo co = failWithL (hang (ptext (sLit "Ill-kinded coercion term:")) 2 (ppr co))
+
+-------------------
+lintType :: OutType -> LintM Kind
+lintType (TyVarTy tv)
+  = do { checkTyVarInScope tv
+       ; return (tyVarKind tv) }
+
+lintType ty@(AppTy t1 t2) 
+  = do { k1 <- lintType t1
+       ; lint_ty_app ty k1 [t2] }
+
+lintType ty@(FunTy t1 t2)
+  = lint_ty_app ty (tyConKind funTyCon) [t1,t2]
+
+lintType ty@(TyConApp tc tys)
+  | tyConHasKind tc
+  = lint_ty_app ty (tyConKind tc) tys
+  | otherwise
+  = failWithL (hang (ptext (sLit "Malformed type:")) 2 (ppr ty))
+
+lintType (ForAllTy tv ty)
+  = do { lintTyBndrKind tv
+       ; addInScopeVar tv (lintType ty) }
+
+lintType ty@(PredTy (ClassP cls tys))
+  = lint_ty_app ty (tyConKind (classTyCon cls)) tys
+
+lintType (PredTy (IParam _ p_ty))
+  = lintType p_ty
+
+lintType ty@(PredTy (EqPred {}))
+  = failWithL (badEq ty)
+
+----------------
+lint_ty_app :: Type -> Kind -> [OutType] -> LintM Kind
+lint_ty_app ty k tys 
+  = do { ks <- mapM lintType tys
+       ; lint_kind_app (ptext (sLit "type") <+> quotes (ppr ty)) k ks }
+                      
+----------------
+check_co_app :: Coercion -> Kind -> [OutType] -> LintM ()
+check_co_app ty k tys 
+  = do { _ <- lint_kind_app (ptext (sLit "coercion") <+> quotes (ppr ty))  
+                            k (map typeKind tys)
+       ; return () }
+                      
+----------------
+lint_kind_app :: SDoc -> Kind -> [Kind] -> LintM Kind
+lint_kind_app doc kfn ks = go kfn ks
+  where
+    fail_msg = vcat [hang (ptext (sLit "Kind application error in")) 2 doc,
+                            nest 2 (ptext (sLit "Function kind =") <+> ppr kfn),
+                            nest 2 (ptext (sLit "Arg kinds =") <+> ppr ks)]
+
+    go kfn []     = return kfn
+    go kfn (k:ks) = case splitKindFunTy_maybe kfn of
+                             Nothing         -> failWithL fail_msg
+                     Just (kfa, kfb) -> do { unless (k `isSubKind` kfa)
+                                                     (addErrL fail_msg)
+                                            ; go kfb ks } 
+--------------
+badEq :: Type -> SDoc
+badEq ty = hang (ptext (sLit "Unexpected equality predicate:"))
+              1 (quotes (ppr ty))
+\end{code}
     
 %************************************************************************
 %*                                                                     *
@@ -619,6 +792,7 @@ data LintLocInfo
   | AnExpr CoreExpr    -- Some expression
   | ImportedUnfolding SrcLoc -- Some imported unfolding (ToDo: say which)
   | TopLevelBindings
+  | InType Type                -- Inside a type
 \end{code}
 
                  
@@ -672,12 +846,16 @@ inCasePat = LintM $ \ loc _ errs -> (Just (is_case_pat loc), errs)
 addInScopeVars :: [Var] -> LintM a -> LintM a
 addInScopeVars vars m
   | null dups
-  = LintM (\ loc subst errs -> unLintM m loc (extendTvInScope subst vars) errs)
+  = LintM (\ loc subst errs -> unLintM m loc (extendTvInScopeList subst vars) errs)
   | otherwise
   = failWithL (dupVars dups)
   where
     (_, dups) = removeDups compare vars 
 
+addInScopeVar :: Var -> LintM a -> LintM a
+addInScopeVar var m
+  = LintM (\ loc subst errs -> unLintM m loc (extendTvInScope subst var) errs)
+
 updateTvSubst :: TvSubst -> LintM a -> LintM a
 updateTvSubst subst' m = 
   LintM (\ loc _ errs -> unLintM m loc subst' errs)
@@ -768,6 +946,8 @@ dumpLoc (ImportedUnfolding locn)
   = (locn, brackets (ptext (sLit "in an imported unfolding")))
 dumpLoc TopLevelBindings
   = (noSrcLoc, empty)
+dumpLoc (InType ty)
+  = (noSrcLoc, text "In the type" <+> quotes (ppr ty))
 
 pp_binders :: [Var] -> SDoc
 pp_binders bs = sep (punctuate comma (map pp_binder bs))
index b5c48fe..84e7ee9 100644 (file)
@@ -1302,6 +1302,13 @@ numericTyKeys =
        , doubleTyConKey
        , floatTyConKey
        ]
+
+kindKeys :: [Unique] 
+kindKeys = [ liftedTypeKindTyConKey
+          , openTypeKindTyConKey
+          , unliftedTypeKindTyConKey
+          , ubxTupleKindTyConKey 
+          , argTypeKindTyConKey ]
 \end{code}
 
 
index 1996e70..597b025 100644 (file)
@@ -1,9 +1,9 @@
-T%
+%
 % (c) The University of Glasgow 2006
 %
 
 \begin{code}
-{-# OPTIONS -fno-warn-incomplete-patterns #-}
+{-# OPTIONS -w #-}
 -- The above warning supression flag is a temporary kludge.
 -- While working on this module you are encouraged to remove it and fix
 -- any warnings in the module. See
@@ -21,7 +21,7 @@ module Coercion (
         -- * Main data type
         Coercion,
  
-        mkCoKind, mkReflCoKind, coVarKind,
+        mkCoKind, mkCoPredTy, coVarKind, coVarKind_maybe,
         coercionKind, coercionKinds, isIdentityCoercion,
 
        -- ** Equality predicates
@@ -30,7 +30,7 @@ module Coercion (
        -- ** Coercion transformations
        mkCoercion,
         mkSymCoercion, mkTransCoercion,
-        mkLeftCoercion, mkRightCoercion, mkRightCoercions,
+        mkLeftCoercion, mkRightCoercion, 
        mkInstCoercion, mkAppCoercion, mkTyConCoercion, mkFunCoercion,
         mkForAllCoercion, mkInstsCoercion, mkUnsafeCoercion,
         mkNewTypeCoercion, mkFamInstCoercion, mkAppsCoercion,
@@ -43,6 +43,9 @@ module Coercion (
         rightCoercionTyCon, instCoercionTyCon, -- needed by TysWiredIn
         csel1CoercionTyCon, csel2CoercionTyCon, cselRCoercionTyCon, 
 
+        -- ** Decomposition
+        decompLR_maybe, decompCsel_maybe, decompInst_maybe,
+
         -- ** Optimisation
        optCoercion,
 
@@ -70,7 +73,9 @@ import Var
 import Name
 import PrelNames
 import Util
+import Control.Monad
 import BasicTypes
+import MonadUtils
 import Outputable
 import FastString
 
@@ -102,15 +107,22 @@ decomposeCo n co
 
 coVarKind :: CoVar -> (Type,Type) 
 -- c :: t1 ~ t2
-coVarKind cv = splitCoVarKind (tyVarKind cv)
+coVarKind cv = case coVarKind_maybe cv of
+                 Just ts -> ts
+                 Nothing -> pprPanic "coVarKind" (ppr cv $$ ppr (tyVarKind cv))
+
+coVarKind_maybe :: CoVar -> Maybe (Type,Type) 
+coVarKind_maybe cv = splitCoKind_maybe (tyVarKind cv)
 
 -- | Take a 'CoercionKind' apart into the two types it relates: see also 'mkCoKind'.
 -- Panics if the argument is not a valid 'CoercionKind'
-splitCoVarKind :: Kind -> (Type, Type)
-splitCoVarKind co | Just co' <- kindView co = splitCoVarKind co'
-splitCoVarKind (PredTy (EqPred ty1 ty2))    = (ty1, ty2)
+splitCoKind_maybe :: Kind -> Maybe (Type, Type)
+splitCoKind_maybe co | Just co' <- kindView co = splitCoKind_maybe co'
+splitCoKind_maybe (PredTy (EqPred ty1 ty2))    = Just (ty1, ty2)
+splitCoKind_maybe _                            = Nothing
 
--- | Makes a 'CoercionKind' from two types: the types whose equality is proven by the relevant 'Coercion'
+-- | Makes a 'CoercionKind' from two types: the types whose equality 
+-- is proven by the relevant 'Coercion'
 mkCoKind :: Type -> Type -> CoercionKind
 mkCoKind ty1 ty2 = PredTy (EqPred ty1 ty2)
 
@@ -118,6 +130,15 @@ mkCoKind ty1 ty2 = PredTy (EqPred ty1 ty2)
 mkCoPredTy :: Type -> Type -> Type -> Type
 mkCoPredTy s t r = ForAllTy (mkWildCoVar (mkCoKind s t)) r
 
+splitCoPredTy_maybe :: Type -> Maybe (Type, Type, Type)
+splitCoPredTy_maybe ty
+  | Just (cv,r) <- splitForAllTy_maybe ty
+  , isCoVar cv
+  , Just (s,t) <- coVarKind_maybe cv
+  = Just (s,t,r)
+  | otherwise
+  = Nothing
+
 -- | Tests whether a type is just a type equality predicate
 isEqPredTy :: Type -> Bool
 isEqPredTy (PredTy pred) = isEqPred pred
@@ -133,10 +154,6 @@ getEqPredTys :: PredType -> (Type,Type)
 getEqPredTys (EqPred ty1 ty2) = (ty1, ty2)
 getEqPredTys other           = pprPanic "getEqPredTys" (ppr other)
 
--- | Create a reflexive 'CoercionKind' that asserts that a type can be coerced to itself
-mkReflCoKind :: Type -> CoercionKind
-mkReflCoKind ty = mkCoKind ty ty
-
 -- | If it is the case that
 --
 -- > c :: (t1 ~ t2)
@@ -146,16 +163,20 @@ coercionKind :: Coercion -> (Type, Type)
 coercionKind ty@(TyVarTy a) | isCoVar a = coVarKind a
                             | otherwise = (ty, ty)
 coercionKind (AppTy ty1 ty2) 
-  = let (t1, t2) = coercionKind ty1
-        (s1, s2) = coercionKind ty2 in
-    (mkAppTy t1 s1, mkAppTy t2 s2)
-coercionKind (TyConApp tc args)
+  = let (s1, t1) = coercionKind ty1
+        (s2, t2) = coercionKind ty2 in
+    (mkAppTy s1 s2, mkAppTy t1 t2)
+coercionKind co@(TyConApp tc args)
   | Just (ar, rule) <- isCoercionTyCon_maybe tc 
     -- CoercionTyCons carry their kinding rule, so we use it here
-  = ASSERT( length args >= ar )        -- Always saturated
-    let (ty1,ty2)    = rule (take ar args)     -- Apply the rule to the right number of args
-       (tys1, tys2) = coercionKinds (drop ar args)
-    in (mkAppTys ty1 tys1, mkAppTys ty2 tys2)
+  = WARN( not (length args >= ar), ppr co )    -- Always saturated
+    (let (ty1,ty2) = runID (rule (return . typeKind)
+                                (return . coercionKind)
+                               False (take ar args))
+                             -- Apply the rule to the right number of args
+                             -- Always succeeds (if term is well-kinded!)
+        (tys1, tys2) = coercionKinds (drop ar args)
+     in (mkAppTys ty1 tys1, mkAppTys ty2 tys2))
 
   | otherwise
   = let (lArgs, rArgs) = coercionKinds args in
@@ -262,97 +283,25 @@ mkSymCoercion :: Coercion -> Coercion
 -- ^ Create a symmetric version of the given 'Coercion' that asserts equality
 -- between the same types but in the other "direction", so a kind of @t1 ~ t2@ 
 -- becomes the kind @t2 ~ t1@.
---
--- This function attempts to simplify the generated 'Coercion' by removing 
--- redundant applications of @sym@. This is done by pushing this new @sym@ 
--- down into the 'Coercion' and exploiting the fact that @sym (sym co) = co@.
-mkSymCoercion co      
-  | Just co' <- coreView co = mkSymCoercion co'
-
-mkSymCoercion (ForAllTy tv ty)  = ForAllTy tv (mkSymCoercion ty)
-mkSymCoercion (AppTy co1 co2)  = AppTy (mkSymCoercion co1) (mkSymCoercion co2)
-mkSymCoercion (FunTy co1 co2)  = FunTy (mkSymCoercion co1) (mkSymCoercion co2)
-
-mkSymCoercion (TyConApp tc cos) 
-  | not (isCoercionTyCon tc) = mkTyConApp tc (map mkSymCoercion cos)
-
-mkSymCoercion (TyConApp tc [co]) 
-  | tc `hasKey` symCoercionTyConKey   = co    -- sym (sym co) --> co
-  | tc `hasKey` leftCoercionTyConKey  = mkLeftCoercion (mkSymCoercion co)
-  | tc `hasKey` rightCoercionTyConKey = mkRightCoercion (mkSymCoercion co)
-
-mkSymCoercion (TyConApp tc [co1,co2]) 
-  | tc `hasKey` transCoercionTyConKey
-     -- sym (co1 `trans` co2) --> (sym co2) `trans (sym co2)
-     -- Note reversal of arguments!
-  = mkTransCoercion (mkSymCoercion co2) (mkSymCoercion co1)
-
-  | tc `hasKey` instCoercionTyConKey
-     -- sym (co @ ty) --> (sym co) @ ty
-     -- Note: sym is not applied to 'ty'
-  = mkInstCoercion (mkSymCoercion co1) co2
-
-mkSymCoercion (TyConApp tc cos)        -- Other coercion tycons, such as those
-  = mkCoercion symCoercionTyCon [TyConApp tc cos]  -- arising from newtypes
-
-mkSymCoercion (TyVarTy tv) 
-  | isCoVar tv = mkCoercion symCoercionTyCon [TyVarTy tv]
-  | otherwise  = TyVarTy tv    -- Reflexive
-
--------------------------------
--- ToDo: we should be cleverer about transitivity
+mkSymCoercion g = mkCoercion symCoercionTyCon [g]
 
 mkTransCoercion :: Coercion -> Coercion -> Coercion
 -- ^ Create a new 'Coercion' by exploiting transitivity on the two given 'Coercion's.
--- 
--- This function attempts to simplify the generated 'Coercion' by exploiting the fact that
--- @sym g `trans` g = id@.
-mkTransCoercion g1 g2  -- sym g `trans` g = id
-  | (t1,_) <- coercionKind g1
-  , (_,t2) <- coercionKind g2
-  , t1 `coreEqType` t2 
-  = t1 
-
-  | otherwise
-  = mkCoercion transCoercionTyCon [g1, g2]
-
-
--------------------------------
--- Smart constructors for left and right
+mkTransCoercion g1 g2 = mkCoercion transCoercionTyCon [g1, g2]
 
 mkLeftCoercion :: Coercion -> Coercion
 -- ^ From an application 'Coercion' build a 'Coercion' that asserts the equality of 
 -- the "functions" on either side of the type equality. So if @c@ has kind @f x ~ g y@ then:
 --
 -- > mkLeftCoercion c :: f ~ g
-mkLeftCoercion co 
-  | Just (co', _) <- splitAppCoercion_maybe co = co'
-  | otherwise = mkCoercion leftCoercionTyCon [co]
+mkLeftCoercion co = mkCoercion leftCoercionTyCon [co]
 
 mkRightCoercion :: Coercion -> Coercion
 -- ^ From an application 'Coercion' build a 'Coercion' that asserts the equality of 
 -- the "arguments" on either side of the type equality. So if @c@ has kind @f x ~ g y@ then:
 --
 -- > mkLeftCoercion c :: x ~ y
-mkRightCoercion  co      
-  | Just (_, co2) <- splitAppCoercion_maybe co = co2
-  | otherwise = mkCoercion rightCoercionTyCon [co]
-
-mkRightCoercions :: Int -> Coercion -> [Coercion]
--- ^ As 'mkRightCoercion', but finds the 'Coercion's available on the right side of @n@
--- nested application 'Coercion's, manufacturing new left or right cooercions as necessary
--- if suffficiently many are not directly available.
-mkRightCoercions n co
-  = go n co []
-  where
-    go n co acc 
-       | n > 0
-       = case splitAppCoercion_maybe co of
-          Just (co1,co2) -> go (n-1) co1 (co2:acc)
-          Nothing        -> go (n-1) (mkCoercion leftCoercionTyCon [co]) (mkCoercion rightCoercionTyCon [co]:acc)
-       | otherwise
-       = acc
-
+mkRightCoercion co = mkCoercion rightCoercionTyCon [co]
 
 mkCsel1Coercion, mkCsel2Coercion, mkCselRCoercion :: Coercion -> Coercion
 mkCsel1Coercion co = mkCoercion csel1CoercionTyCon [co]
@@ -363,72 +312,12 @@ mkCselRCoercion co = mkCoercion cselRCoercionTyCon [co]
 mkInstCoercion :: Coercion -> Type -> Coercion
 -- ^ Instantiates a 'Coercion' with a 'Type' argument. If possible, it immediately performs
 -- the resulting beta-reduction, otherwise it creates a suspended instantiation.
-mkInstCoercion co ty
-  | Just (tv,co') <- splitForAllTy_maybe co
-  = substTyWith [tv] [ty] co'  -- (forall a.co) @ ty  -->  co[ty/a]
-  | otherwise
-  = mkCoercion instCoercionTyCon  [co, ty]
+mkInstCoercion co ty = mkCoercion instCoercionTyCon  [co, ty]
 
 mkInstsCoercion :: Coercion -> [Type] -> Coercion
 -- ^ As 'mkInstCoercion', but instantiates the coercion with a number of type arguments, left-to-right
 mkInstsCoercion co tys = foldl mkInstCoercion co tys
 
-{-
-splitSymCoercion_maybe :: Coercion -> Maybe Coercion
-splitSymCoercion_maybe (TyConApp tc [co]) = 
-  if tc `hasKey` symCoercionTyConKey
-  then Just co
-  else Nothing
-splitSymCoercion_maybe co = Nothing
--}
-
-splitAppCoercion_maybe :: Coercion -> Maybe (Coercion, Coercion)
--- ^ Splits a coercion application, being careful *not* to split @left c@ etc.
--- This is because those are really syntactic constructs, not applications
-splitAppCoercion_maybe co  | Just co' <- coreView co = splitAppCoercion_maybe co'
-splitAppCoercion_maybe (FunTy ty1 ty2)   = Just (TyConApp funTyCon [ty1], ty2)
-splitAppCoercion_maybe (AppTy ty1 ty2)   = Just (ty1, ty2)
-splitAppCoercion_maybe (TyConApp tc tys) 
-   | not (isCoercionTyCon tc)
-   = case snocView tys of
-       Just (tys', ty') -> Just (TyConApp tc tys', ty')
-       Nothing          -> Nothing
-splitAppCoercion_maybe _ = Nothing
-
-{-
-splitTransCoercion_maybe :: Coercion -> Maybe (Coercion, Coercion)
-splitTransCoercion_maybe (TyConApp tc [ty1, ty2]) 
- = if tc `hasKey` transCoercionTyConKey then
-       Just (ty1, ty2)
-   else
-       Nothing
-splitTransCoercion_maybe other = Nothing
-
-splitInstCoercion_maybe :: Coercion -> Maybe (Coercion, Type)
-splitInstCoercion_maybe (TyConApp tc [ty1, ty2])
- = if tc `hasKey` instCoercionTyConKey then
-       Just (ty1, ty2)
-    else
-       Nothing
-splitInstCoercion_maybe other = Nothing
-
-splitLeftCoercion_maybe :: Coercion -> Maybe Coercion
-splitLeftCoercion_maybe (TyConApp tc [co])
- = if tc `hasKey` leftCoercionTyConKey then
-       Just co
-   else
-       Nothing
-splitLeftCoercion_maybe other = Nothing
-
-splitRightCoercion_maybe :: Coercion -> Maybe Coercion
-splitRightCoercion_maybe (TyConApp tc [co])
- = if tc `hasKey` rightCoercionTyConKey then
-       Just co
-   else
-       Nothing
-splitRightCoercion_maybe other = Nothing
--}
-
 -- | Manufacture a coercion from this air. Needless to say, this is not usually safe,
 -- but it is used when we know we are dealing with bottom, which is one case in which 
 -- it is safe.  This is also used implement the @unsafeCoerce#@ primitive.
@@ -449,8 +338,12 @@ mkNewTypeCoercion name tycon tvs rhs_ty
   where
     co_con_arity = length tvs
 
-    rule args = ASSERT( co_con_arity == length args )
-               (TyConApp tycon args, substTyWith tvs args rhs_ty)
+    rule :: CoTyConKindChecker
+    rule kc_ty kc_co checking args 
+      = do { ks <- mapM kc_ty args
+           ; unless (not checking || kindAppOk (tyConKind tycon) ks) 
+                    (fail "Argument kind mis-match")
+           ; return (TyConApp tycon args, substTyWith tvs args rhs_ty) }
 
 -- | Create a coercion identifying a @data@, @newtype@ or @type@ representation type
 -- and its family instance.  It has the form @Co tvs :: F ts ~ R tvs@, where @Co@ is 
@@ -466,9 +359,22 @@ mkFamInstCoercion name tvs family instTys rep_tycon
   = mkCoercionTyCon name coArity rule
   where
     coArity = length tvs
-    rule args = (substTyWith tvs args $                     -- with sigma = [tys/tvs],
-                  TyConApp family instTys,          --       sigma (F ts)
-                TyConApp rep_tycon args)            --   ~ R tys
+
+    rule :: CoTyConKindChecker
+    rule kc_ty kc_co checking args 
+      = do { ks <- mapM kc_ty args
+           ; unless (not checking  || kindAppOk (tyConKind rep_tycon) ks)
+                    (fail "Argument kind mis-match")
+           ; return (substTyWith tvs args $         -- with sigma = [tys/tvs],
+                    TyConApp family instTys         --       sigma (F ts)
+                   , TyConApp rep_tycon args) }     --   ~ R tys
+
+kindAppOk :: Kind -> [Kind] -> Bool
+kindAppOk kfn [] = True
+kindAppOk kfn (k:ks) 
+  = case splitKindFunTy_maybe kfn of
+      Just (kfa, kfb) | k `isSubKind` kfa -> kindAppOk kfb ks
+      _other                              -> False
 \end{code}
 
 
@@ -497,92 +403,103 @@ symCoercionTyCon, transCoercionTyCon, leftCoercionTyCon,
   rightCoercionTyCon, instCoercionTyCon, unsafeCoercionTyCon,
   csel1CoercionTyCon, csel2CoercionTyCon, cselRCoercionTyCon :: TyCon
 
-symCoercionTyCon = 
-  mkCoercionTyCon symCoercionTyConName 1 flipCoercionKindOf
+symCoercionTyCon 
+  = mkCoercionTyCon symCoercionTyConName 1 kc_sym
   where
-    flipCoercionKindOf (co:rest) = ASSERT( null rest ) (ty2, ty1)
-       where
-         (ty1, ty2) = coercionKind co
+    kc_sym :: CoTyConKindChecker
+    kc_sym kc_ty kc_co _ (co:_) 
+      = do { (ty1,ty2) <- kc_co co
+           ; return (ty2,ty1) }
 
-transCoercionTyCon = 
-  mkCoercionTyCon transCoercionTyConName 2 composeCoercionKindsOf
+transCoercionTyCon 
+  = mkCoercionTyCon transCoercionTyConName 2 kc_trans
   where
-    composeCoercionKindsOf (co1:co2:rest)
-      = ASSERT( null rest )
-        WARN( not (r1 `coreEqType` a2),
-              text "Strange! Type mismatch in trans coercion, probably a bug"
-              $$
-             _err_stuff )
-        (a1, r2)
-      where
-        (a1, r1) = coercionKind co1
-        (a2, r2) = coercionKind co2 
-
-        _err_stuff = vcat [ text "co1:" <+> ppr co1
-                          , text "co1 kind left:"  <+> ppr a1
-                          , text "co1 kind right:" <+> ppr r1
-                          , text "co2:" <+> ppr co2
-                          , text "co2 kind left:"  <+> ppr a2
-                          , text "co2 kind right:" <+> ppr r2 ]
+    kc_trans :: CoTyConKindChecker
+    kc_trans kc_ty kc_co checking (co1:co2:_)
+      = do { (a1, r1) <- kc_co co1
+           ; (a2, r2) <- kc_co co2 
+          ; unless (not checking || (r1 `coreEqType` a2))
+                    (fail "Trans coercion mis-match")
+           ; return (a1, r2) }
 
 ---------------------------------------------------
-leftCoercionTyCon  = mkCoercionTyCon leftCoercionTyConName  1 (fst . decompLR)
-rightCoercionTyCon = mkCoercionTyCon rightCoercionTyConName 1 (snd . decompLR)
+leftCoercionTyCon  = mkCoercionTyCon leftCoercionTyConName  1 (kcLR_help fst)
+rightCoercionTyCon = mkCoercionTyCon rightCoercionTyConName 1 (kcLR_help snd)
+
+kcLR_help :: (forall a. (a,a)->a) -> CoTyConKindChecker
+kcLR_help select kc_ty kc_co _checking (co : _)
+  = do { (ty1, ty2)  <- kc_co co
+       ; case decompLR_maybe ty1 ty2 of
+           Nothing  -> fail "decompLR" 
+           Just res -> return (select res) }
 
-decompLR :: [Type] -> ((Type,Type), (Type,Type))
+decompLR_maybe :: Type -> Type -> Maybe ((Type,Type), (Type,Type))
 -- Helper for left and right.  Finds coercion kind of its input and
 -- returns the left and right projections of the coercion...
 --
 -- if c :: t1 s1 ~ t2 s2 then splitCoercionKindOf c = ((t1, t2), (s1, s2))
-decompLR (co : rest)
-  | (ty1, ty2) <- coercionKind co
-  , Just (ty_fun1, ty_arg1) <- splitAppTy_maybe ty1
+decompLR_maybe ty1 ty2
+  | Just (ty_fun1, ty_arg1) <- splitAppTy_maybe ty1
   , Just (ty_fun2, ty_arg2) <- splitAppTy_maybe ty2
-  = ASSERT( null rest) 
-    ((ty_fun1, ty_fun2),(ty_arg1, ty_arg2))
-decompLR cos 
-  = pprPanic "Coercion.decompLR" 
-             (ppr cos $$ vcat (map (pprEqPred .coercionKind) cos))
+  = Just ((ty_fun1, ty_fun2),(ty_arg1, ty_arg2))
+decompLR_maybe _ _ = Nothing
 
 ---------------------------------------------------
 instCoercionTyCon 
-  =  mkCoercionTyCon instCoercionTyConName 2 instCoercionKind
+  =  mkCoercionTyCon instCoercionTyConName 2 kcInst_help
   where
-    instantiateCo t s =
-      let Just (tv, ty) = splitForAllTy_maybe t in
-      substTyWith [tv] [s] ty
+    kcInst_help :: CoTyConKindChecker
+    kcInst_help kc_ty kc_co checking (co : ty : _)
+      = do { (t1,t2) <- kc_co co
+           ; k <- kc_ty ty
+           ; case decompInst_maybe t1 t2 of
+               Nothing -> fail "decompInst"
+               Just ((tv1,tv2), (ty1,ty2)) -> do
+           { unless (not checking || (k `isSubKind` tyVarKind tv1))
+                    (fail "Coercion instantation kind mis-match")
+           ; return (substTyWith [tv1] [ty] ty1,
+                     substTyWith [tv2] [ty] ty2) } }
+
+decompInst_maybe :: Type -> Type -> Maybe ((TyVar,TyVar), (Type,Type))
+decompInst_maybe ty1 ty2
+  | Just (tv1,r1) <- splitForAllTy_maybe ty1
+  , Just (tv2,r2) <- splitForAllTy_maybe ty2
+  = Just ((tv1,tv2), (r1,r2))
 
-    instCoercionKind (co1:ty:rest) = ASSERT( null rest )
-                                    (instantiateCo t1 ty, instantiateCo t2 ty)
-      where (t1, t2) = coercionKind co1
 
 ---------------------------------------------------
 unsafeCoercionTyCon 
-  = mkCoercionTyCon unsafeCoercionTyConName 2 unsafeCoercionKind
+  = mkCoercionTyCon unsafeCoercionTyConName 2 kc_unsafe
   where
-   unsafeCoercionKind (ty1:ty2:rest) = ASSERT( null rest ) (ty1,ty2) 
+   kc_unsafe kc_ty kc_co _checking (ty1:ty2:_) 
+    = do { k1 <- kc_ty ty1
+         ; k2 <- kc_ty ty2
+         ; return (ty1,ty2) }
         
 ---------------------------------------------------
 -- The csel* family
+
+csel1CoercionTyCon = mkCoercionTyCon csel1CoercionTyConName 1 (kcCsel_help fstOf3)
+csel2CoercionTyCon = mkCoercionTyCon csel2CoercionTyConName 1 (kcCsel_help sndOf3)
+cselRCoercionTyCon = mkCoercionTyCon cselRCoercionTyConName 1 (kcCsel_help thirdOf3) 
+
+kcCsel_help :: (forall a. (a,a,a) -> a) -> CoTyConKindChecker
+kcCsel_help select kc_ty kc_co _checking (co : rest)
+  = do { (ty1,ty2) <- kc_co co
+       ; case decompCsel_maybe ty1 ty2 of
+           Nothing  -> fail "decompCsel"
+           Just res -> return (select res) }
+
+decompCsel_maybe :: Type -> Type -> Maybe ((Type,Type), (Type,Type), (Type,Type))
 --   If         co :: (s1~t1 => r1) ~ (s2~t2 => r2)
--- Then   csel1 co :: s1 ~ s2
---        csel2 co :: t1 ~ t2
---        cselR co :: r1 ~ r2
-
-csel1CoercionTyCon = mkCoercionTyCon csel1CoercionTyConName 1 (fstOf3   . decompCsel)
-csel2CoercionTyCon = mkCoercionTyCon csel2CoercionTyConName 1 (sndOf3   . decompCsel)
-cselRCoercionTyCon = mkCoercionTyCon cselRCoercionTyConName 1 (thirdOf3 . decompCsel)
-
-decompCsel :: [Coercion] -> ((Type,Type), (Type,Type), (Type,Type))
-decompCsel (co : rest)
-  | (ty1,ty2) <- coercionKind co
-  , Just (cv1, r1) <- splitForAllTy_maybe ty1
-  , Just (cv2, r2) <- splitForAllTy_maybe ty2
-  , (s1,t1) <- ASSERT( isCoVar cv1) coVarKind cv1
-  , (s2,t2) <- ASSERT( isCoVar cv1) coVarKind cv2
-  = ASSERT( null rest )
-    ((s1,s2), (t1,t2), (r1,r2))
-decompCsel other = pprPanic "decompCsel" (ppr other)
+-- Then   csel1 co ::            s1 ~ s2
+--        csel2 co ::           t1 ~ t2
+--        cselR co ::           r1 ~ r2
+decompCsel_maybe ty1 ty2
+  | Just (s1, t1, r1) <- splitCoPredTy_maybe ty1
+  , Just (s2, t2, r2) <- splitCoPredTy_maybe ty2
+  = Just ((s1,s2), (t1,t2), (r1,r2))
+decompCsel_maybe _ _ = Nothing
 
 fstOf3   :: (a,b,c) -> a    
 sndOf3   :: (a,b,c) -> b    
@@ -729,7 +646,7 @@ mkAppTyCoI _   IdCo _   IdCo = IdCo
 mkAppTyCoI ty1 coi1 ty2 coi2 =
        ACo $ AppTy (fromCoI coi1 ty1) (fromCoI coi2 ty2)
 
--- | Smart constructor for function-'Coercion's on 'CoercionI', see also 'mkFunCoercion'
+
 mkFunTyCoI :: Type -> CoercionI -> Type -> CoercionI -> CoercionI
 mkFunTyCoI _   IdCo _   IdCo = IdCo
 mkFunTyCoI ty1 coi1 ty2 coi2 =
@@ -772,182 +689,209 @@ mkEqPredCoI _   (ACo co1) ty2 coi2      = ACo $ PredTy $ EqPred co1 (fromCoI coi
 %************************************************************************
 
 \begin{code}
-optCoercion :: Coercion -> Coercion
-optCoercion co
-  = ASSERT2( coercionKind co `eq` coercionKind result, 
-             ppr co $$ ppr result $$ ppr (coercionKind co) $$ ppr (coercionKind result) )
-    result
+type NormalCo = Coercion
+  -- Invariants: 
+  --  * For trans coercions (co1 `trans` co2)
+  --       co1 is not a trans, and neither co1 nor co2 is identity
+  --  * If the coercion is the identity, it has no CoVars of CoTyCons in it (just types)
+
+type NormalNonIdCo = NormalCo  -- Extra invariant: not the identity
+
+optCoercion :: Coercion -> NormalCo
+optCoercion co = opt_co False co
+
+opt_co :: Bool        -- True <=> return (sym co)
+       -> Coercion
+       -> NormalCo
+opt_co = opt_co'
+-- opt_co sym co = pprTrace "opt_co {" (ppr sym <+> ppr co) $
+--                     co1 `seq` 
+--                pprTrace "opt_co done }" (ppr co1) 
+--               WARN( not same_co_kind, ppr co  <+> dcolon <+> pprEqPred (s1,t1) 
+--                                     $$ ppr co1 <+> dcolon <+> pprEqPred (s2,t2) )
+--                co1
+--  where
+--    co1 = opt_co' sym co
+--    same_co_kind = s1 `coreEqType` s2 && t1 `coreEqType` t2
+--    (s,t) = coercionKind co
+--    (s1,t1) | sym = (t,s)
+--            | otherwise = (s,t)
+--    (s2,t2) = coercionKind co1
+
+opt_co' sym (AppTy ty1 ty2)          = mkAppTy (opt_co sym ty1) (opt_co sym ty2)
+opt_co' sym (FunTy ty1 ty2)          = FunTy (opt_co sym ty1) (opt_co sym ty2)
+opt_co' sym (PredTy (ClassP cls tys)) = PredTy (ClassP cls (map (opt_co sym) tys))
+opt_co' sym (PredTy (IParam n ty))    = PredTy (IParam n (opt_co sym ty))
+
+opt_co' sym co@(TyVarTy tv)
+  | not (isCoVar tv)     = co   -- Identity; does not mention a CoVar
+  | ty1 `coreEqType` ty2 = ty1 -- Identity; ..ditto..
+  | not sym              = co
+  | otherwise            = mkSymCoercion co
+  where
+    (ty1,ty2) = coVarKind tv
+
+opt_co' sym (ForAllTy tv cor) 
+  | isCoVar tv = mkCoPredTy (opt_co sym co1) (opt_co sym co2) (opt_co sym cor)
+  | otherwise  = ForAllTy tv (opt_co sym cor)
+  where
+    (co1,co2) = coVarKind tv
+
+opt_co' sym (TyConApp tc cos)
+  | isCoercionTyCon tc
+  = foldl mkAppTy opt_co_tc 
+          (map (opt_co sym) (drop arity cos))
+  | otherwise
+  = TyConApp tc (map (opt_co sym) cos)
+  where
+    arity = tyConArity tc
+    opt_co_tc :: NormalCo
+    opt_co_tc = opt_co_tc_app sym tc (take arity cos)
+
+--------
+opt_co_tc_app :: Bool -> TyCon -> [Type] -> NormalCo
+-- Used for CoercionTyCons only
+opt_co_tc_app sym tc cos
+  | tc `hasKey` symCoercionTyConKey
+  = opt_co (not sym) co1
+
+  | tc `hasKey` transCoercionTyConKey
+  = if sym then opt_trans opt_co2 opt_co1
+           else opt_trans opt_co1 opt_co2
+
+  | tc `hasKey` leftCoercionTyConKey
+  , Just (co1, _) <- splitAppTy_maybe opt_co1
+  = co1
+
+  | tc `hasKey` rightCoercionTyConKey
+  , Just (_, co2) <- splitAppTy_maybe opt_co1
+  = co2
+
+  | tc `hasKey` csel1CoercionTyConKey
+  , Just (s1,_,_) <- splitCoPredTy_maybe opt_co1
+  = s1
+
+  | tc `hasKey` csel2CoercionTyConKey
+  , Just (_,s2,_) <- splitCoPredTy_maybe opt_co1
+  = s2
+
+  | tc `hasKey` cselRCoercionTyConKey
+  , Just (_,_,r) <- splitCoPredTy_maybe opt_co1
+  = r
+
+  | tc `hasKey` instCoercionTyConKey
+  , Just (tv, co'') <- splitForAllTy_maybe opt_co1
+  , let ty = co2
+  = substTyWith [tv] [ty] co''
+
+  | otherwise    -- Do not push sym inside top-level axioms
+                 -- e.g. if g is a top-level axiom
+                 --   g a : F a ~ a
+                 -- Then (sym (g ty)) /= g (sym ty) !!
+  = if sym then mkSymCoercion the_co 
+           else the_co
+  where
+    the_co = TyConApp tc cos
+    (co1 : cos1) = cos
+    (co2 : _)    = cos1
+    opt_co1 = opt_co sym co1
+    opt_co2 = opt_co sym co2
+
+-------------
+opt_trans :: NormalCo -> NormalCo -> NormalCo
+opt_trans co1 co2
+  | isIdNormCo co1 = co2
+  | otherwise      = opt_trans1 co1 co2
+
+opt_trans1 :: NormalNonIdCo -> NormalCo -> NormalCo
+-- First arg is not the identity
+opt_trans1 co1 co2
+  | isIdNormCo co2 = co1
+  | otherwise      = opt_trans2 co1 co2
+
+opt_trans2 :: NormalNonIdCo -> NormalNonIdCo -> NormalCo
+-- Neither arg is the identity
+opt_trans2 (TyConApp tc [co1a,co1b]) co2
+  | tc `hasKey` transCoercionTyConKey
+  = opt_trans1 co1a (opt_trans2 co1b co2)
+
+opt_trans2 co1 co2 
+  | Just co <- opt_trans_rule co1 co2
+  = co
+
+opt_trans2 co1 (TyConApp tc [co2a,co2b])
+  | tc `hasKey` transCoercionTyConKey
+  , Just co1_2a <- opt_trans_rule co1 co2a
+  = if isIdNormCo co1_2a
+    then co2b
+    else opt_trans2 co1_2a co2b
+
+opt_trans2 co1 co2
+  = mkTransCoercion co1 co2
+
+------
+opt_trans_rule :: NormalNonIdCo -> NormalNonIdCo -> Maybe NormalCo
+opt_trans_rule (TyConApp tc [co1]) co2
+  | tc `hasKey` symCoercionTyConKey
+  , co1 `coreEqType` co2
+  , (_,ty2) <- coercionKind co2
+  = Just ty2
+
+opt_trans_rule co1 (TyConApp tc [co2])
+  | tc `hasKey` symCoercionTyConKey
+  , co1 `coreEqType` co2
+  , (ty1,_) <- coercionKind co1
+  = Just ty1
+
+opt_trans_rule (TyConApp tc1 [co1,ty1]) (TyConApp tc2 [co2,ty2])
+  | tc1 `hasKey` instCoercionTyConKey
+  , tc1 == tc2
+  , ty1 `coreEqType` ty2
+  = Just (mkInstCoercion (opt_trans2 co1 co2) ty1) 
+
+opt_trans_rule (TyConApp tc1 cos1) (TyConApp tc2 cos2)
+  | not (isCoercionTyCon tc1) || 
+    getUnique tc1 `elem` [ leftCoercionTyConKey, rightCoercionTyConKey
+                         , csel1CoercionTyConKey, csel2CoercionTyConKey
+                        , cselRCoercionTyConKey ]      --Yuk!
+  , tc1 == tc2                  -- Works for left,right, and csel* family
+                        -- BUT NOT equality axioms 
+                        -- E.g.        (g Int) `trans` (g Bool)
+                        --        /= g (Int . Bool)
+  = Just (TyConApp tc1 (zipWith opt_trans cos1 cos2))
+
+opt_trans_rule co1 co2
+  | Just (co1a, co1b) <- splitAppTy_maybe co1
+  , Just (co2a, co2b) <- splitAppTy_maybe co2
+  = Just (mkAppTy (opt_trans co1a co2a) (opt_trans co1b co2b))
+
+  | Just (s1,t1,r1) <- splitCoPredTy_maybe co1
+  , Just (s2,t2,r2) <- splitCoPredTy_maybe co1
+  = Just (mkCoPredTy (opt_trans s1 s2)
+                     (opt_trans t1 t2)
+                     (opt_trans r1 r2))
+
+  | Just (tv1,r1) <- splitForAllTy_maybe co1
+  , Just (tv2,r2) <- splitForAllTy_maybe co2
+  , not (isCoVar tv1)               -- Both have same kind
+  , let r2' = substTyWith [tv2] [TyVarTy tv1] r2
+  = Just (ForAllTy tv1 (opt_trans2 r1 r2'))
+
+opt_trans_rule _ _ = Nothing
+
+  
+-------------
+isIdNormCo :: NormalCo -> Bool
+-- Cheap identity test: look for coercions with no coercion variables at all
+-- So it'll return False for (sym g `trans` g)
+isIdNormCo ty = go ty
   where
-        (s1,t1) `eq` (s2,t2) = s1 `coreEqType` s2 && t1 `coreEqType` t2
-
-        (result,_,_) = go co
-                         -- optimized, changed?, identity?
-        go :: Coercion -> ( Coercion, Bool, Bool )
-        -- traverse coercion term bottom up and return
-        --
-        --    1) equivalent coercion, in optimized form
-        --
-        --    2) whether the output coercion differs from
-        --       the input coercion
-        --
-        --    3) whether the coercion is an identity coercion
-        --
-        -- Performs the following optimizations:
-        --
-        --      sym id          >->     id
-        --      trans id co     >->     co
-        --      trans co id     >->     co
-        --      sym (sym co)    >->     co
-       --      trans g (sym g) >->     id
-       --      trans (sym g) g >->     id
-        --
-        go ty@(TyVarTy a) | isCoVar a = let (ty1,ty2) = coercionKind ty
-                                        in (ty, False, ty1 `coreEqType` ty2)
-                          | otherwise = (ty, False, True)
-        go ty@(AppTy ty1 ty2)
-          = let (ty1', chan1, id1) = go ty1
-                (ty2', chan2, id2) = go ty2
-            in if chan1 || chan2
-                 then (AppTy ty1' ty2', True,  id1 && id2)
-                 else (ty             , False, id1 && id2)
-        go ty@(TyConApp tc args)
-          | tc == symCoercionTyCon, (ty1:tys) <- args
-          = goSym ty ty1 tys
-          | tc == transCoercionTyCon, [ty1,ty2] <- args
-          = goTrans ty ty1 ty2
-          | tc == leftCoercionTyCon, [ty1] <- args
-         = goLeft ty ty1
-          | tc == rightCoercionTyCon, [ty1] <- args
-         = goRight ty ty1
-         | tc == instCoercionTyCon, [ty1,ty2] <- args
-         = goInst ty ty1 ty2
-         | not (isCoercionTyCon tc)
-          = let (args', chans, ids) = mapAndUnzip3 go args
-            in  if or chans
-                  then (TyConApp tc args', True , and ids)
-                  else (ty               , False, and ids) 
-          | otherwise
-          = (ty, False, False)
-        go ty@(FunTy ty1 ty2)
-          = let (ty1',chan1,id1) = go ty1
-                (ty2',chan2,id2) = go ty2
-            in  if chan1 || chan2
-                  then (FunTy ty1' ty2', True , id1 && id2)
-                  else (ty             , False, id1 && id2)
-        go ty@(ForAllTy tv ty1)
-          = let (ty1', chan1, id1) = go ty1
-            in if chan1
-                 then (ForAllTy tv ty1', True , id1)
-                 else (ty              , False, id1)
-        go ty@(PredTy (EqPred ty1 ty2))
-          = let (ty1', chan1, id1) = go ty1
-                (ty2', chan2, id2) = go ty2
-            in if chan1 || chan2
-                 then (PredTy (EqPred ty1' ty2'), True , id1 && id2)
-                 else (ty                       , False, id1 && id2)
-        go ty@(PredTy (ClassP cl args))
-          = let (args', chans, ids) = mapAndUnzip3 go args
-            in  if or chans
-                  then (PredTy (ClassP cl args'), True , and ids)
-                  else (ty                      , False, and ids)
-        go ty@(PredTy (IParam name ty1))
-          = let (ty1', chan1, id1) = go ty1
-            in  if chan1
-                  then (PredTy (IParam name ty1'), True , id1)
-                  else (ty                       , False, id1)
-
-        goSym :: Coercion -> Coercion -> [Coercion] -> ( Coercion, Bool, Bool )
-       --
-        -- pushes the sym constructor inwards, if possible
-       --
-       --   takes original coercion term
-       --                  first argument 
-        --                  rest of arguments
-        goSym ty ty1 tys  
-          = case mkSymCoercion ty1 of
-             (TyConApp tc _ ) | tc == symCoercionTyCon
-                               -> let (tys',chans',ids) = mapAndUnzip3 go (ty1:tys)
-                                  in  if or chans'
-                                        then (TyConApp symCoercionTyCon tys', True , and ids)
-                                        else (ty                            , False, and ids)
-              ty1'             -> let (ty',_   ,id') = go (mkAppsCoercion ty1' tys)
-                                  in  (ty',True,id')
-
-
-        goRight :: Coercion -> Coercion -> ( Coercion, Bool, Bool )
-       --
-        -- reduces the right constructor, if possible
-       --
-       --   takes original coercion term
-       --                  argument 
-       --
-        goRight ty ty1  
-          = case mkRightCoercion ty1 of
-             (TyConApp tc _ ) | tc == rightCoercionTyCon
-                               -> let (ty1',chan1,id1) = go ty1
-                                  in  if chan1
-                                        then (TyConApp rightCoercionTyCon [ty1'], True , id1)
-                                        else (ty                                , False, id1)
-              ty1'             -> let (ty',_   ,id') = go ty1'
-                                  in  (ty',True,id')
-
-        goLeft :: Coercion -> Coercion -> ( Coercion, Bool, Bool )
-       --
-        -- reduces the left constructor, if possible
-       --
-       --   takes original coercion term
-       --                  argument 
-       --
-        goLeft ty ty1  
-          = case mkLeftCoercion ty1 of
-             (TyConApp tc _ ) | tc == leftCoercionTyCon
-                               -> let (ty1',chan1,id1) = go ty1
-                                  in  if chan1
-                                        then (TyConApp leftCoercionTyCon [ty1'], True , id1)
-                                        else (ty                                , False, id1)
-              ty1'             -> let (ty',_   ,id') = go ty1'
-                                  in  (ty',True,id')
-
-        goInst :: Coercion -> Coercion -> Coercion -> ( Coercion, Bool, Bool )
-       --
-        -- reduces the inst constructor, if possible
-       --
-       --   takes original coercion term
-       --                  coercion argument 
-       --                  type argument
-       --
-        goInst ty ty1 ty2
-          = case mkInstCoercion ty1 ty2 of
-             (TyConApp tc _ ) | tc == instCoercionTyCon
-                               -> let (ty1',chan1,id1) = go ty1
-                                  in  if chan1
-                                        then (TyConApp instCoercionTyCon [ty1',ty2], True , id1)
-                                        else (ty                                   , False, id1)
-              ty1'             -> let (ty',_   ,id') = go ty1'
-                                  in  (ty',True,id')
-
-        goTrans :: Coercion -> Coercion -> Coercion -> ( Coercion, Bool, Bool )
-       --
-        --      trans id co     >->     co
-        --      trans co id     >->     co
-       --      trans g (sym g) >->     id
-       --      trans (sym g) g >->     id
-       --
-        goTrans ty ty1 ty2
-         | id1
-          = (ty2', True, id2)
-          | id2
-          = (ty1', True, False)
-          | chan1 || chan2
-         = (TyConApp transCoercionTyCon [ty1',ty2'], True , False)
-         | Just ty' <- mty'
-          = (ty', True, True)
-          | otherwise
-          = (ty, False, False)
-          where (ty1', chan1, id1) = go ty1
-                (ty2', chan2, id2) = go ty2
-               mty' = case mkTransCoercion ty1' ty2'
-                         of (TyConApp tc _) | tc == transCoercionTyCon
-                                             -> Nothing
-                            ty'              -> Just ty' 
+    go (TyVarTy tv)           = not (isCoVar tv)
+    go (AppTy t1 t2)          = go t1 && go t2
+    go (FunTy t1 t2)          = go t1 && go t2
+    go (ForAllTy tv ty)        = go (tyVarKind tv) && go ty
+    go (TyConApp tc tys)       = not (isCoercionTyCon tc) && all go tys
+    go (PredTy (IParam _ ty))  = go ty
+    go (PredTy (ClassP _ tys)) = all go tys
+    go (PredTy (EqPred t1 t2)) = go t1 && go t2
 \end{code}  
index 6f8803c..43fb524 100644 (file)
@@ -8,7 +8,7 @@ The @TyCon@ datatype
 \begin{code}
 module TyCon(
         -- * Main TyCon data types
-       TyCon, FieldLabel,
+       TyCon, FieldLabel, CoTyConKindChecker,
 
        AlgTyConRhs(..), visibleDataCons, 
         TyConParent(..), 
@@ -36,9 +36,9 @@ module TyCon(
         isPrimTyCon,
         isTupleTyCon, isUnboxedTupleTyCon, isBoxedTupleTyCon, 
         isSynTyCon, isClosedSynTyCon, isOpenSynTyCon,
-        isSuperKindTyCon,
+        isSuperKindTyCon, isDecomposableTyCon,
         isCoercionTyCon, isCoercionTyCon_maybe,
-        isForeignTyCon, isAnyTyCon,
+        isForeignTyCon, isAnyTyCon, tyConHasKind,
 
        isInjectiveTyCon,
        isDataTyCon, isProductTyCon, isEnumerationTyCon, 
@@ -124,7 +124,7 @@ data TyCon
     FunTyCon {
        tyConUnique :: Unique,
        tyConName   :: Name,
-       tyConKind   :: Kind,
+       tc_kind   :: Kind,
        tyConArity  :: Arity
     }
 
@@ -133,7 +133,7 @@ data TyCon
   | AlgTyCon {         
        tyConUnique :: Unique,
        tyConName   :: Name,
-       tyConKind   :: Kind,
+       tc_kind   :: Kind,
        tyConArity  :: Arity,
 
        tyConTyVars :: [TyVar],         -- ^ The type variables used in the type constructor.
@@ -171,7 +171,7 @@ data TyCon
   | TupleTyCon {
        tyConUnique :: Unique,
        tyConName   :: Name,
-       tyConKind   :: Kind,
+       tc_kind   :: Kind,
        tyConArity  :: Arity,
        tyConBoxed  :: Boxity,
        tyConTyVars :: [TyVar],
@@ -183,7 +183,7 @@ data TyCon
   | SynTyCon {
        tyConUnique  :: Unique,
        tyConName    :: Name,
-       tyConKind    :: Kind,
+       tc_kind    :: Kind,
        tyConArity   :: Arity,
 
        tyConTyVars  :: [TyVar],        -- Bound tyvars
@@ -199,14 +199,14 @@ data TyCon
   | PrimTyCon {                        
        tyConUnique   :: Unique,
        tyConName     :: Name,
-       tyConKind     :: Kind,
+       tc_kind     :: Kind,
        tyConArity    :: Arity,                 -- SLPJ Oct06: I'm not sure what the significance
                                                --             of the arity of a primtycon is!
 
        primTyConRep  :: PrimRep,               -- ^ Many primitive tycons are unboxed, but some are
                                                        --   boxed (represented by pointers). This 'PrimRep' holds
                                                --   that information.
-                                               -- Only relevant if tyConKind = *
+                                               -- Only relevant if tc_kind = *
 
        isUnLifted   :: Bool,                   -- ^ Most primitive tycons are unlifted (may not contain bottom)
                                                --   but foreign-imported ones may be lifted
@@ -216,18 +216,14 @@ data TyCon
     }
 
   -- | Type coercions, such as @(~)@, @sym@, @trans@, @left@ and @right@.
-  -- INVARIANT: coercions are always fully applied
+  -- INVARIANT: Coercion TyCons are always fully applied
+  --           But note that a CoercionTyCon can be over-saturated in a type.
+  --           E.g.  (sym g1) Int  will be represented as (TyConApp sym [g1,Int])
   | CoercionTyCon {    
        tyConUnique :: Unique,
         tyConName   :: Name,
        tyConArity  :: Arity,
-       coKindFun   :: [Type] -> (Type,Type)
-               -- ^ Function that when given a list of the type arguments to the 'TyCon'
-               -- constructs the types that the resulting coercion relates.
-               --
-               -- INVARIANT: 'coKindFun' is always applied to exactly 'tyConArity' args
-               -- E.g. for @trans (c1 :: ta=tb) (c2 :: tb=tc)@, the 'coKindFun' returns 
-               --      the kind as a pair of types: @(ta, tc)@
+       coKindFun   :: CoTyConKindChecker
     }
 
   -- | Any types.  Like tuples, this is a potentially-infinite family of TyCons
@@ -239,7 +235,7 @@ data TyCon
   | AnyTyCon {
        tyConUnique  :: Unique,
        tyConName    :: Name,
-       tyConKind    :: Kind    -- Never = *; that is done via PrimTyCon
+       tc_kind    :: Kind      -- Never = *; that is done via PrimTyCon
                                -- See Note [Any types] in TysPrim
     }
 
@@ -254,6 +250,23 @@ data TyCon
         tyConName   :: Name
     }
 
+type CoTyConKindChecker = forall m. Monad m => CoTyConKindCheckerFun m
+
+type CoTyConKindCheckerFun m 
+  =    (Type -> m Kind)                -- Kind checker for types
+    -> (Type -> m (Type,Type)) -- and for coercions
+    -> Bool                    -- True => apply consistency checks
+    -> [Type]                  -- Exactly right number of args
+    -> m (Type, Type)          -- Kind of this application
+
+               -- ^ Function that when given a list of the type arguments to the 'TyCon'
+               -- constructs the types that the resulting coercion relates.
+               -- Returns Nothing if ill-kinded.
+               --
+               -- INVARIANT: 'coKindFun' is always applied to exactly 'tyConArity' args
+               -- E.g. for @trans (c1 :: ta=tb) (c2 :: tb=tc)@, the 'coKindFun' returns 
+               --      the kind as a pair of types: @(ta, tc)@
+
 -- | Names of the fields in an algebraic record type
 type FieldLabel = Name
 
@@ -578,7 +591,7 @@ mkFunTyCon name kind
   = FunTyCon { 
        tyConUnique = nameUnique name,
        tyConName   = name,
-       tyConKind   = kind,
+       tc_kind   = kind,
        tyConArity  = 2
     }
 
@@ -598,7 +611,7 @@ mkAlgTyCon name kind tyvars stupid rhs parent is_rec gen_info gadt_syn
   = AlgTyCon { 
        tyConName        = name,
        tyConUnique      = nameUnique name,
-       tyConKind        = kind,
+       tc_kind  = kind,
        tyConArity       = length tyvars,
        tyConTyVars      = tyvars,
        algTcStupidTheta = stupid,
@@ -626,7 +639,7 @@ mkTupleTyCon name kind arity tyvars con boxed gen_info
   = TupleTyCon {
        tyConUnique = nameUnique name,
        tyConName = name,
-       tyConKind = kind,
+       tc_kind = kind,
        tyConArity = arity,
        tyConBoxed = boxed,
        tyConTyVars = tyvars,
@@ -647,7 +660,7 @@ mkForeignTyCon name ext_name kind arity
   = PrimTyCon {
        tyConName    = name,
        tyConUnique  = nameUnique name,
-       tyConKind    = kind,
+       tc_kind    = kind,
        tyConArity   = arity,
        primTyConRep = PtrRep, -- they all do
        isUnLifted   = False,
@@ -675,7 +688,7 @@ mkPrimTyCon' name kind arity rep is_unlifted
   = PrimTyCon {
        tyConName    = name,
        tyConUnique  = nameUnique name,
-       tyConKind    = kind,
+       tc_kind    = kind,
        tyConArity   = arity,
        primTyConRep = rep,
        isUnLifted   = is_unlifted,
@@ -688,7 +701,7 @@ mkSynTyCon name kind tyvars rhs parent
   = SynTyCon { 
        tyConName = name,
        tyConUnique = nameUnique name,
-       tyConKind = kind,
+       tc_kind = kind,
        tyConArity = length tyvars,
        tyConTyVars = tyvars,
        synTcRhs = rhs,
@@ -696,19 +709,27 @@ mkSynTyCon name kind tyvars rhs parent
     }
 
 -- | Create a coercion 'TyCon'
-mkCoercionTyCon :: Name -> Arity -> ([Type] -> (Type,Type)) -> TyCon
-mkCoercionTyCon name arity kindRule
+mkCoercionTyCon :: Name -> Arity 
+                -> CoTyConKindChecker
+                -> TyCon
+mkCoercionTyCon name arity rule_fn
   = CoercionTyCon {
-        tyConName = name,
+        tyConName   = name,
         tyConUnique = nameUnique name,
-        tyConArity = arity,
-        coKindFun = kindRule
+        tyConArity  = arity,
+#ifdef DEBUG
+        coKindFun   = \ ty co fail args -> 
+                      ASSERT2( length args == arity, ppr name )
+                      rule_fn ty co fail args
+#else
+       coKindFun   = rule_fn
+#endif
     }
 
 mkAnyTyCon :: Name -> Kind -> TyCon
 mkAnyTyCon name kind 
-  = AnyTyCon { tyConName = name,
-               tyConKind = kind,
+  = AnyTyCon {  tyConName = name,
+               tc_kind = kind,
                tyConUnique = nameUnique name }
 
 -- | Create a super-kind 'TyCon'
@@ -778,6 +799,11 @@ isNewTyCon :: TyCon -> Bool
 isNewTyCon (AlgTyCon {algTcRhs = NewTyCon {}}) = True
 isNewTyCon _                                   = False
 
+tyConHasKind :: TyCon -> Bool
+tyConHasKind (SuperKindTyCon {}) = False
+tyConHasKind (CoercionTyCon {})  = False
+tyConHasKind _                   = True
+
 -- | Take a 'TyCon' apart into the 'TyVar's it scopes over, the 'Type' it expands
 -- into, and (possibly) a coercion from the representation type to the @newtype@.
 -- Returns @Nothing@ if this is not possible.
@@ -823,6 +849,13 @@ isClosedSynTyCon tycon = isSynTyCon tycon && not (isOpenTyCon tycon)
 isOpenSynTyCon :: TyCon -> Bool
 isOpenSynTyCon tycon = isSynTyCon tycon && isOpenTyCon tycon
 
+isDecomposableTyCon :: TyCon -> Bool
+-- True iff we can deocmpose (T a b c) into ((T a b) c)
+-- Specifically NOT true of synonyms (open and otherwise) and coercions
+isDecomposableTyCon (SynTyCon      {}) = False
+isDecomposableTyCon (CoercionTyCon {}) = False
+isDecomposableTyCon _other             = True
+
 -- | Is this an algebraic 'TyCon' declared with the GADT syntax?
 isGadtSyntaxTyCon :: TyCon -> Bool
 isGadtSyntaxTyCon (AlgTyCon { algTcGadtSyntax = res }) = res
@@ -938,7 +971,7 @@ isAnyTyCon _              = False
 -- | Attempt to pull a 'TyCon' apart into the arity and 'coKindFun' of
 -- a coercion 'TyCon'. Returns @Nothing@ if the 'TyCon' is not of the
 -- appropriate kind
-isCoercionTyCon_maybe :: TyCon -> Maybe (Arity, [Type] -> (Type,Type))
+isCoercionTyCon_maybe :: Monad m => TyCon -> Maybe (Arity, CoTyConKindCheckerFun m)
 isCoercionTyCon_maybe (CoercionTyCon {tyConArity = ar, coKindFun = rule}) 
   = Just (ar, rule)
 isCoercionTyCon_maybe _ = Nothing
@@ -1023,6 +1056,15 @@ tyConHasGenerics (AlgTyCon {hasGenerics = hg})   = hg
 tyConHasGenerics (TupleTyCon {hasGenerics = hg}) = hg
 tyConHasGenerics _                               = False        -- Synonyms
 
+tyConKind :: TyCon -> Kind
+tyConKind (FunTyCon   { tc_kind = k }) = k
+tyConKind (AlgTyCon   { tc_kind = k }) = k
+tyConKind (TupleTyCon { tc_kind = k }) = k
+tyConKind (SynTyCon   { tc_kind = k }) = k
+tyConKind (PrimTyCon  { tc_kind = k }) = k
+tyConKind (AnyTyCon   { tc_kind = k }) = k
+tyConKind tc                           = pprPanic "tyConKind" (ppr tc)
+
 -- | As 'tyConDataCons_maybe', but returns the empty list of constructors if no constructors
 -- could be found
 tyConDataCons :: TyCon -> [DataCon]
index 8dfe475..630340a 100644 (file)
@@ -64,7 +64,7 @@ module Type (
         Kind, SimpleKind, KindVar,
         
         -- ** Deconstructing Kinds 
-        kindFunResult, splitKindFunTys, splitKindFunTysN,
+        kindFunResult, splitKindFunTys, splitKindFunTysN, splitKindFunTy_maybe,
 
         -- ** Common Kinds and SuperKinds
         liftedTypeKind, unliftedTypeKind, openTypeKind,
@@ -122,7 +122,7 @@ module Type (
        emptyTvSubstEnv, emptyTvSubst,
        
        mkTvSubst, mkOpenTvSubst, zipOpenTvSubst, zipTopTvSubst, mkTopTvSubst, notElemTvSubst,
-       getTvSubstEnv, setTvSubstEnv, getTvInScope, extendTvInScope,
+       getTvSubstEnv, setTvSubstEnv, getTvInScope, extendTvInScope, extendTvInScopeList,
        extendTvSubst, extendTvSubstList, isInScope, composeTvSubst, zipTyEnv,
         isEmptyTvSubst,
 
@@ -403,7 +403,7 @@ repSplitAppTy_maybe :: Type -> Maybe (Type,Type)
 repSplitAppTy_maybe (FunTy ty1 ty2)   = Just (TyConApp funTyCon [ty1], ty2)
 repSplitAppTy_maybe (AppTy ty1 ty2)   = Just (ty1, ty2)
 repSplitAppTy_maybe (TyConApp tc tys) 
-  | not (isOpenSynTyCon tc) || length tys > tyConArity tc 
+  | isDecomposableTyCon tc || length tys > tyConArity tc 
   = case snocView tys of       -- never create unsaturated type family apps
       Just (tys', ty') -> Just (TyConApp tc tys', ty')
       Nothing         -> Nothing
@@ -427,9 +427,9 @@ splitAppTys ty = split ty ty []
     split _       (AppTy ty arg)        args = split ty ty (arg:args)
     split _       (TyConApp tc tc_args) args
       = let -- keep type families saturated
-            n | isOpenSynTyCon tc = tyConArity tc
-              | otherwise         = 0
-            (tc_args1, tc_args2)  = splitAt n tc_args
+            n | isDecomposableTyCon tc = tyConArity tc
+              | otherwise              = 0
+            (tc_args1, tc_args2) = splitAt n tc_args
         in
         (TyConApp tc tc_args1, tc_args2 ++ args)
     split _       (FunTy ty1 ty2)       args = ASSERT( null args )
@@ -1433,8 +1433,11 @@ notElemTvSubst tv (TvSubst _ env) = not (tv `elemVarEnv` env)
 setTvSubstEnv :: TvSubst -> TvSubstEnv -> TvSubst
 setTvSubstEnv (TvSubst in_scope _) env = TvSubst in_scope env
 
-extendTvInScope :: TvSubst -> [Var] -> TvSubst
-extendTvInScope (TvSubst in_scope env) vars = TvSubst (extendInScopeSetList in_scope vars) env
+extendTvInScope :: TvSubst -> Var -> TvSubst
+extendTvInScope (TvSubst in_scope env) var = TvSubst (extendInScopeSet in_scope var) env
+
+extendTvInScopeList :: TvSubst -> [Var] -> TvSubst
+extendTvInScopeList (TvSubst in_scope env) vars = TvSubst (extendInScopeSetList in_scope vars) env
 
 extendTvSubst :: TvSubst -> TyVar -> Type -> TvSubst
 extendTvSubst (TvSubst in_scope env) tv ty = TvSubst in_scope (extendVarEnv env tv ty)
@@ -1720,6 +1723,9 @@ kindFunResult k = funResultTy k
 splitKindFunTys :: Kind -> ([Kind],Kind)
 splitKindFunTys k = splitFunTys k
 
+splitKindFunTy_maybe :: Kind -> Maybe (Kind,Kind)
+splitKindFunTy_maybe = splitFunTy_maybe
+
 -- | Essentially 'splitFunTysN' on kinds
 splitKindFunTysN :: Int -> Kind -> ([Kind],Kind)
 splitKindFunTysN k = splitFunTysN k