Type families: bug fixes
[ghc-hetmet.git] / compiler / typecheck / TcTyFuns.lhs
index 126a029..84453bc 100644 (file)
@@ -3,16 +3,16 @@ normalisation and entailment checking of equality constraints.
 
 \begin{code}
 module TcTyFuns (
-       tcNormaliseFamInst,
+  -- type normalisation wrt to toplevel equalities only
+  tcNormaliseFamInst,
 
-       normaliseGivenEqs, normaliseGivenDicts, 
-       normaliseWantedEqs, normaliseWantedDicts,
-       solveWantedEqs,
-       substEqInDictInsts,
-       
-        -- errors
-        eqInstMisMatch, misMatchMsg,
-  ) where
+  -- instance normalisation wrt to equalities
+  tcReduceEqs,
+
+  -- errors
+  misMatchMsg, failWithMisMatch,
+
+) where
 
 
 #include "HsVersions.h"
@@ -30,23 +30,26 @@ import Type
 import TypeRep         ( Type(..) )
 import TyCon
 import HsSyn
+import Id
 import VarEnv
+import VarSet
 import Var
 import Name
 import Bag
 import Outputable
 import SrcLoc  ( Located(..) )
 import Maybes
+import FastString
 
 -- standard
 import Data.List
-import Control.Monad (liftM)
+import Control.Monad
 \end{code}
 
 
 %************************************************************************
 %*                                                                     *
-               Normalisation of types
+               Normalisation of types wrt toplevel equality schemata
 %*                                                                     *
 %************************************************************************
 
@@ -76,7 +79,7 @@ tcUnfoldSynFamInst (TyConApp tycon tys)
                                                    mkTyConApp coe_tc tys')
              where
                tys'   = rep_tys ++ restTys
-               coe_tc = expectJust "TcTyFun.tcUnfoldSynFamInst" 
+               coe_tc = expectJust "TcTyFuns.tcUnfoldSynFamInst" 
                                    (tyConFamilyCoercion_maybe rep_tc)
        }
     where
@@ -92,38 +95,12 @@ possible (ie, we treat family instances as a TRS).  Also zonk meta variables.
        then   co : ty ~ ty'
 
 \begin{code}
+-- |Normalise the given type as far as possible with toplevel equalities.
+-- This results in a coercion witnessing the type equality, in addition to the
+-- normalised type.
+--
 tcNormaliseFamInst :: TcType -> TcM (CoercionI, TcType)
 tcNormaliseFamInst = tcGenericNormaliseFamInst tcUnfoldSynFamInst
-
-tcNormaliseFamInstPred :: TcPredType -> TcM (CoercionI, TcPredType)
-tcNormaliseFamInstPred = tcGenericNormaliseFamInstPred tcUnfoldSynFamInst
-\end{code}
-
-Normalise a type relative to the rewrite rule implied by one EqInst or an
-already unpacked EqInst (ie, an equality rewrite rule).
-
-\begin{code}
--- Rewrite by EqInst
-tcEqInstNormaliseFamInst :: Inst -> TcType -> TcM (CoercionI, Type)
-tcEqInstNormaliseFamInst inst = 
-  ASSERT( isEqInst inst )
-  tcEqRuleNormaliseFamInst (pat, rhs, co)
-  where
-    pat = eqInstLeftTy  inst
-    rhs = eqInstRightTy inst
-    co  = eqInstType    inst
-
--- Rewrite by equality rewrite rule
-tcEqRuleNormaliseFamInst :: (TcType,   -- rewrite rule lhs
-                             TcType,   -- rewrite rule rhs
-                             TcType)   -- coercion witnessing the rewrite rule
-                          -> TcType    -- type to rewrite
-                          -> TcM (CoercionI, Type)
-tcEqRuleNormaliseFamInst (pat, rhs, co) ty
-  = tcGenericNormaliseFamInst matchEqRule ty
-  where
-    matchEqRule sty | pat `tcEqType` sty = return $ Just (rhs, co)
-                    | otherwise          = return $ Nothing
 \end{code}
 
 Generic normalisation of 'Type's and 'PredType's; ie, walk the type term and
@@ -178,10 +155,6 @@ tcGenericNormaliseFamInst fun (ForAllTy tyvar ty1)
   = do         { (coi,nty1) <- tcGenericNormaliseFamInst fun ty1
        ; return (mkForAllTyCoI tyvar coi, mkForAllTy tyvar nty1)
        }
-tcGenericNormaliseFamInst fun (NoteTy note ty1)
-  = do { (coi,nty1) <- tcGenericNormaliseFamInst fun ty1
-       ; return (mkNoteTyCoI note coi, NoteTy note nty1)
-       }
 tcGenericNormaliseFamInst fun ty@(TyVarTy tv)
   | isTcTyVar tv
   = do { traceTc (text "tcGenericNormaliseFamInst" <+> ppr ty)
@@ -226,752 +199,1016 @@ tcGenericNormaliseFamInstPred fun (EqPred ty1 ty2)
 
 %************************************************************************
 %*                                                                     *
-\section{Normalisation of equality constraints}
+               Normalisation of instances wrt to equalities
 %*                                                                     *
 %************************************************************************
 
-Note [Inconsistencies in equality constraints]
-~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
-We guarantee that we raise an error if we discover any inconsistencies (i.e.,
-equalities that if presented to the unifer in TcUnify would result in an
-error) during normalisation of wanted constraints.  This is especially so that
-we don't solve wanted constraints under an inconsistent given set.  In
-particular, we don't want to permit signatures, such as
-
-  bad :: (Int ~ Bool => Int) -> a -> a
-
 \begin{code}
-normaliseGivenEqs :: [Inst] -> TcM ([Inst], TcM ())
-normaliseGivenEqs givens
- = do { traceTc (text "normaliseGivenEqs <-" <+> ppr givens)
-      ; (result, deSkolem) <- 
-          rewriteToFixedPoint (Just ("(SkolemOccurs)", skolemOccurs))
-           [ ("(ZONK)",    dontRerun $ zonkInsts)
-           , ("(TRIVIAL)", dontRerun $ trivialRule)
-           , ("(DECOMP)",  decompRule)
-           , ("(TOP)",     topRule)
-           , ("(SUBST)",   substRule)                   -- incl. occurs check
-            ] givens
-      ; traceTc (text "normaliseGivenEqs ->" <+> ppr result)
-      ; return (result, deSkolem)
-      }
-\end{code}
-
-\begin{code}
-normaliseWantedEqs :: [Inst] -> TcM [Inst]
-normaliseWantedEqs insts 
-  = do { traceTc (text "normaliseWantedEqs <-" <+> ppr insts)
-       ; result <- liftM fst $ rewriteToFixedPoint Nothing
-                    [ ("(ZONK)",    dontRerun $ zonkInsts)
-                    , ("(TRIVIAL)", dontRerun $ trivialRule)
-                    , ("(DECOMP)",  decompRule)
-                    , ("(TOP)",     topRule)
-                    , ("(UNIFY)",   unifyMetaRule)      -- incl. occurs check
-                    , ("(SUBST)",   substRule)          -- incl. occurs check
-                     ] insts
-       ; traceTc (text "normaliseWantedEqs ->" <+> ppr result)
-       ; return result
+tcReduceEqs :: [Inst]             -- locals
+            -> [Inst]             -- wanteds
+            -> TcM ([Inst],       -- normalised locals (w/o equalities)
+                    [Inst],       -- normalised wanteds (including equalities)
+                    TcDictBinds,  -- bindings for all simplified dictionaries
+                    Bool)         -- whether any flexibles where instantiated
+tcReduceEqs locals wanteds
+  = do { let (local_eqs  , local_dicts)   = partition isEqInst locals
+             (wanteds_eqs, wanteds_dicts) = partition isEqInst wanteds
+       ; eqCfg1 <- normaliseEqs (local_eqs ++ wanteds_eqs)
+       ; eqCfg2 <- normaliseDicts False local_dicts
+       ; eqCfg3 <- normaliseDicts True  wanteds_dicts
+       ; eqCfg <- propagateEqs (eqCfg1 `unionEqConfig` eqCfg2 
+                                       `unionEqConfig` eqCfg3)
+       ; finaliseEqsAndDicts eqCfg
        }
 \end{code}
 
 
 %************************************************************************
 %*                                                                     *
-\section{Solving of wanted constraints with respect to a given set}
+               Equality Configurations
 %*                                                                     *
 %************************************************************************
 
-The set of given equalities must have been normalised already.
+We maintain normalised equalities together with the skolems introduced as
+intermediates during flattening of equalities as well as 
+
+!!!TODO: We probably now can do without the skolem set.  It's not used during
+finalisation in the current code.
 
 \begin{code}
-solveWantedEqs :: [Inst]        -- givens
-            -> [Inst]          -- wanteds
-            -> TcM [Inst]      -- irreducible wanteds
-solveWantedEqs givens wanteds 
-  = do { traceTc $ text "solveWantedEqs <-" <+> ppr wanteds <+> text "with" <+> 
-                   ppr givens
-       ; result <- liftM fst $ rewriteToFixedPoint Nothing
-                     [ ("(ZONK)",    dontRerun $ zonkInsts)
-                     , ("(TRIVIAL)", dontRerun $ trivialRule)
-                     , ("(DECOMP)",  decompRule)
-                     , ("(TOP)",     topRule)
-                     , ("(GIVEN)",   substGivens givens) -- incl. occurs check
-                     , ("(UNIFY)",   unifyMetaRule)      -- incl. occurs check
-                     ] wanteds
-       ; traceTc (text "solveWantedEqs ->" <+> ppr result)
-       ; return result
-       }
-  where
-    -- Use `substInst' with every given on all the wanteds.
-    substGivens :: [Inst] -> [Inst] -> TcM ([Inst], Bool)               
-    substGivens []     wanteds = return (wanteds, False)
-    substGivens (g:gs) wanteds
-      = do { (wanteds1, changed1) <- substGivens gs wanteds
-          ; (wanteds2, changed2) <- substInst g wanteds1
-          ; return (wanteds2, changed1 || changed2)
-          }
+-- |Configuration of normalised equalities used during solving.
+--
+data EqConfig = EqConfig { eqs     :: [RewriteInst]     -- all equalities
+                         , locals  :: [Inst]            -- given dicts
+                         , wanteds :: [Inst]            -- wanted dicts
+                         , binds   :: TcDictBinds       -- bindings
+                         , skolems :: TyVarSet          -- flattening skolems
+                         }
+
+addSkolems :: EqConfig -> TyVarSet -> EqConfig
+addSkolems eqCfg newSkolems 
+  = eqCfg {skolems = skolems eqCfg `unionVarSet` newSkolems}
+
+addEq :: EqConfig -> RewriteInst -> EqConfig
+addEq eqCfg eq = eqCfg {eqs = eq : eqs eqCfg}
+
+unionEqConfig :: EqConfig -> EqConfig -> EqConfig
+unionEqConfig eqc1 eqc2 = EqConfig 
+                          { eqs     = eqs eqc1 ++ eqs eqc2
+                          , locals  = locals eqc1 ++ locals eqc2
+                          , wanteds = wanteds eqc1 ++ wanteds eqc2
+                          , binds   = binds eqc1 `unionBags` binds eqc2
+                          , skolems = skolems eqc1 `unionVarSet` skolems eqc2
+                          }
+
+emptyEqConfig :: EqConfig
+emptyEqConfig = EqConfig
+                { eqs     = []
+                , locals  = []
+                , wanteds = []
+                , binds   = emptyBag
+                , skolems = emptyVarSet
+                }
+
+instance Outputable EqConfig where
+  ppr (EqConfig {eqs = eqs, locals = locals, wanteds = wanteds, binds = binds})
+    = vcat [ppr eqs, ppr locals, ppr wanteds, ppr binds]
 \end{code}
 
-
-%************************************************************************
-%*                                                                     *
-\section{Normalisation of non-equality dictionaries}
-%*                                                                     *
-%************************************************************************
+The set of operations on an equality configuration.  We obtain the initialise
+configuration by normalisation ('normaliseEqs'), solve the equalities by
+propagation ('propagateEqs'), and eventually finalise the configuration when
+no further propoagation is possible.
 
 \begin{code}
-normaliseGivenDicts, normaliseWantedDicts
-       :: [Inst]               -- given equations
-       -> [Inst]               -- dictionaries
-       -> TcM ([Inst],TcDictBinds)
-
-normaliseGivenDicts  eqs dicts = normalise_dicts eqs dicts False
-normaliseWantedDicts eqs dicts = normalise_dicts eqs dicts True
-
-normalise_dicts
-       :: [Inst]       -- given equations
-       -> [Inst]       -- dictionaries
-       -> Bool         -- True <=> the dicts are wanted 
-                       -- Fals <=> they are given
-       -> TcM ([Inst],TcDictBinds)
-normalise_dicts given_eqs dicts is_wanted
-  = do { traceTc $ text "normalise???Dicts <-" <+> ppr dicts <+> 
-                    text "with" <+> ppr given_eqs
-       ; (dicts0, binds0)  <- normaliseInsts is_wanted dicts
-       ; (dicts1, binds1)  <- substEqInDictInsts given_eqs dicts0
-       ; let binds01 = binds0 `unionBags` binds1
-       ; if isEmptyBag binds1
-         then return (dicts1, binds01)
-         else do { (dicts2, binds2) <- normaliseGivenDicts given_eqs dicts1
-                 ; return (dicts2, binds01 `unionBags` binds2) } }
+-- |Turn a set of equalities into an equality configuration for solving.
+--
+-- Precondition: The Insts are zonked.
+--
+normaliseEqs :: [Inst] -> TcM EqConfig
+normaliseEqs eqs 
+  = do { ASSERTM2( allM isValidWantedEqInst eqs, ppr eqs )
+       ; traceTc $ ptext (sLit "Entering normaliseEqs")
+
+       ; (eqss, skolemss) <- mapAndUnzipM normEqInst eqs
+       ; return $ emptyEqConfig { eqs = concat eqss
+                                , skolems = unionVarSets skolemss 
+                                }
+       }
+
+-- |Flatten the type arguments of all dictionaries, returning the result as a 
+-- equality configuration.  The dictionaries go into the 'wanted' component if 
+-- the second argument is 'True'.
+--
+-- Precondition: The Insts are zonked.
+--
+normaliseDicts :: Bool -> [Inst] -> TcM EqConfig
+normaliseDicts isWanted insts
+  = do { traceTc $ ptext (sLit "Entering normaliseDicts") <+>
+                   ptext (if isWanted then sLit "[Wanted]" else sLit "[Local]")
+       ; (insts', eqss, bindss, skolemss) <- mapAndUnzip4M (normDict isWanted) 
+                                                           insts
+       ; return $ emptyEqConfig { eqs     = concat eqss
+                                , locals  = if isWanted then [] else insts'
+                                , wanteds = if isWanted then insts' else []
+                                , binds   = unionManyBags bindss
+                                , skolems = unionVarSets skolemss
+                                }
+       }
+
+-- |Solves the equalities as far as possible by applying propagation rules.
+--
+propagateEqs :: EqConfig -> TcM EqConfig
+propagateEqs eqCfg@(EqConfig {eqs = todoEqs}) 
+  = do { traceTc $ hang (ptext (sLit "Entering propagateEqs:"))
+                     4 (ppr eqCfg)
+
+       ; propagate todoEqs (eqCfg {eqs = []})
+       }
+
+-- |Finalise a set of equalities and associated dictionaries after
+-- propagation.  The returned Boolean value is `True' iff any flexible
+-- variables, except those introduced by flattening (i.e., those in the
+-- `skolems' component of the argument) where instantiated. The first returned
+-- set of instances are the locals (without equalities) and the second set are
+-- all residual wanteds, including equalities. 
+--
+-- Remove all identity dictinary bindings (i.e., those whose source and target
+-- dictionary are the same).  This is important for termination, as
+-- TcSimplify.reduceContext takes the presence of dictionary bindings as an
+-- indicator that there was some improvement.
+--
+finaliseEqsAndDicts :: EqConfig 
+                    -> TcM ([Inst], [Inst], TcDictBinds, Bool)
+finaliseEqsAndDicts (EqConfig { eqs     = eqs
+                              , locals  = locals
+                              , wanteds = wanteds
+                              , binds   = binds
+                              })
+  = do { traceTc $ ptext (sLit "finaliseEqsAndDicts")
+       ; (eqs', subst_binds, locals', wanteds') <- substitute eqs locals wanteds
+       ; (eqs'', improved) <- instantiateAndExtract eqs'
+       ; final_binds <- filterM nonTrivialDictBind $
+                          bagToList (subst_binds `unionBags` binds)
+
+       ; ASSERTM2( allM isValidWantedEqInst eqs'', ppr eqs'' )
+       ; return (locals', eqs'' ++ wanteds', listToBag final_binds, improved)
+       }
+  where
+    nonTrivialDictBind (L _ (VarBind { var_id = ide1
+                                     , var_rhs = L _ (HsWrap _ (HsVar ide2))}))
+      = do { ty1 <- zonkTcType (idType ide1)
+           ; ty2 <- zonkTcType (idType ide2)
+           ; return $ not (ty1 `tcEqType` ty2)
+           }
+    nonTrivialDictBind _ = return True
 \end{code}
 
 
 %************************************************************************
 %*                                                                     *
-\section{Normalisation rules and iterative rule application}
+               Normalisation of equalities
 %*                                                                     *
 %************************************************************************
 
-We have three kinds of normalising rewrite rules:
-
-(1) Normalisation rules that rewrite a set of insts and return a flag indicating
-    whether any changes occurred during rewriting that necessitate re-running
-    the current rule set.
+A normal equality is a properly oriented equality with associated coercion
+that contains at most one family equality (in its left-hand side) is oriented
+such that it may be used as a reqrite rule.  It has one of the following two 
+forms:
 
-(2) Precondition rules that rewrite a set of insts and return a monadic action
-    that reverts the effect of preconditioning.
+(1) co :: F t1..tn ~ t  (family equalities)
+(2) co :: x ~ t         (variable equalities)
 
-(3) Idempotent normalisation rules that never require re-running the rule set. 
+Variable equalities fall again in two classes:
 
-\begin{code}
-type RewriteRule     = [Inst] -> TcM ([Inst], Bool)   -- rewrite, maybe re-run
-type PrecondRule     = [Inst] -> TcM ([Inst], TcM ()) -- rewrite, revertable
-type IdemRewriteRule = [Inst] -> TcM [Inst]           -- rewrite, don't re-run
+(2a) co :: x ~ t, where t is *not* a variable, or
+(2b) co :: x ~ y, where x > y.
 
-type NamedRule       = (String, RewriteRule)          -- rule with description
-type NamedPreRule    = (String, PrecondRule)          -- precond with desc
-\end{code}
+The types t, t1, ..., tn may not contain any occurrences of synonym
+families.  Moreover, in Forms (2) & (3), the left-hand side may not occur in
+the right-hand side, and the relation x > y is an arbitrary, but total order
+on type variables
 
-Template lifting idempotent rules to full rules (which can be put into a rule
-set).
+!!!TODO: We may need to keep track of swapping for error messages (and to
+re-orient on finilisation).
 
 \begin{code}
-dontRerun :: IdemRewriteRule -> RewriteRule
-dontRerun rule insts = liftM addFalse $ rule insts
+data RewriteInst
+  = RewriteVar  -- Form (2) above
+    { rwi_var     :: TyVar    -- may be rigid or flexible
+    , rwi_right   :: TcType   -- contains no synonym family applications
+    , rwi_co      :: EqInstCo -- the wanted or given coercion
+    , rwi_loc     :: InstLoc
+    , rwi_name    :: Name     -- no semantic significance (cf. TcRnTypes.EqInst)
+    , rwi_swapped :: Bool     -- swapped orientation of original EqInst
+    }
+  | RewriteFam  -- Forms (1) above
+    { rwi_fam     :: TyCon    -- synonym family tycon
+    , rwi_args    :: [Type]   -- contain no synonym family applications
+    , rwi_right   :: TcType   -- contains no synonym family applications
+    , rwi_co      :: EqInstCo -- the wanted or given coercion
+    , rwi_loc     :: InstLoc
+    , rwi_name    :: Name     -- no semantic significance (cf. TcRnTypes.EqInst)
+    , rwi_swapped :: Bool     -- swapped orientation of original EqInst
+    }
+
+isWantedRewriteInst :: RewriteInst -> Bool
+isWantedRewriteInst = isWantedCo . rwi_co
+
+rewriteInstToInst :: RewriteInst -> TcM Inst
+rewriteInstToInst eq@(RewriteVar {rwi_var = tv})
+  = deriveEqInst eq (mkTyVarTy tv) (rwi_right eq) (rwi_co eq)
+rewriteInstToInst eq@(RewriteFam {rwi_fam = fam, rwi_args = args})
+  = deriveEqInst eq (mkTyConApp fam args) (rwi_right eq) (rwi_co eq)
+
+-- Derive an EqInst based from a RewriteInst, possibly swapping the types
+-- around. 
+--
+deriveEqInst :: RewriteInst -> TcType -> TcType -> EqInstCo -> TcM Inst
+deriveEqInst rewrite ty1 ty2 co
+  = do { co_adjusted <- if not swapped then return co 
+                                       else mkSymEqInstCo co (ty2, ty1)
+       ; return $ EqInst
+                  { tci_left  = left
+                  , tci_right = right
+                  , tci_co    = co_adjusted
+                  , tci_loc   = rwi_loc rewrite
+                  , tci_name  = rwi_name rewrite
+                  }
+       }
   where
-    addFalse x = (x, False)
+    swapped       = rwi_swapped rewrite
+    (left, right) = if not swapped then (ty1, ty2) else (ty2, ty1)
+
+instance Outputable RewriteInst where
+  ppr (RewriteFam {rwi_fam = fam, rwi_args = args, rwi_right = rhs, rwi_co =co})
+    = hsep [ pprEqInstCo co <+> text "::" 
+           , ppr (mkTyConApp fam args)
+           , text "~>"
+           , ppr rhs
+           ]
+  ppr (RewriteVar {rwi_var = tv, rwi_right = rhs, rwi_co =co})
+    = hsep [ pprEqInstCo co <+> text "::" 
+           , ppr tv
+           , text "~>"
+           , ppr rhs
+           ]
+
+pprEqInstCo :: EqInstCo -> SDoc
+pprEqInstCo (Left cotv) = ptext (sLit "Wanted") <+> ppr cotv
+pprEqInstCo (Right co)  = ptext (sLit "Local") <+> ppr co
 \end{code}
 
-The following function applies a set of rewrite rules until a fixed point is
-reached; i.e., none of the `RewriteRule's require re-running the rule set.
-Optionally, there may be a pre-conditing rule that is applied before any other
-rules are applied and before the rule set is re-run.
+The following functions turn an arbitrary equality into a set of normal
+equalities.  This implements the WFlat and LFlat rules of the paper in one
+sweep.  However, we use flexible variables for both locals and wanteds, and
+avoid to carry around the unflattening substitution \Sigma (for locals) by
+already updating the skolems for locals with the family application that they
+represent - i.e., they will turn into that family application on the next
+zonking (which only happens after finalisation).
+
+In a corresponding manner, normDict normalises class dictionaries by
+extracting any synonym family applications and generation appropriate normal
+equalities. 
 
-The result is the set of rewritten (i.e., normalised) insts and, in case of a
-pre-conditing rule, a monadic action that reverts the effects of
-pre-conditioning - specifically, this is removing introduced skolems.
+Whenever we encounter a loopy equality (of the form a ~ T .. (F ...a...) ...),
+we drop that equality and raise an error if it is a wanted or a warning if it
+is a local.
 
 \begin{code}
-rewriteToFixedPoint :: Maybe NamedPreRule   -- optional preconditioning rule
-                    -> [NamedRule]          -- rule set
-                    -> [Inst]               -- insts to rewrite
-                    -> TcM ([Inst], TcM ())
-rewriteToFixedPoint precondRule rules insts
-  = completeRewrite (return ()) precondRule insts
+normEqInst :: Inst -> TcM ([RewriteInst], TyVarSet)
+-- Normalise one equality.
+normEqInst inst
+  = ASSERT( isEqInst inst )
+    go ty1 ty2 (eqInstCoercion inst)
   where
-    completeRewrite :: TcM () -> Maybe NamedPreRule -> [Inst] 
-                    -> TcM ([Inst], TcM ())
-    completeRewrite dePrecond (Just (precondName, precond)) insts
-      = do { traceTc $ text precondName <+> text " <- " <+> ppr insts
-           ; (insts', dePrecond') <- precond insts
-           ; traceTc $ text precondName <+> text " -> " <+> ppr insts'
-           ; tryRules (dePrecond >> dePrecond') rules insts'
-           }
-    completeRewrite dePrecond Nothing insts
-      = tryRules dePrecond rules insts
-
-    tryRules dePrecond _                    []    = return ([]   , dePrecond)
-    tryRules dePrecond []                   insts = return (insts, dePrecond)
-    tryRules dePrecond ((name, rule):rules) insts 
-      = do { traceTc $ text name <+> text " <- " <+> ppr insts
-           ; (insts', rerun) <- rule insts
-           ; traceTc $ text name <+> text " -> " <+> ppr insts'
-          ; if rerun then completeRewrite dePrecond precondRule insts'
-                     else tryRules dePrecond rules insts'
-           }
-\end{code}
+    (ty1, ty2) = eqInstTys inst
 
+      -- look through synonyms
+    go ty1 ty2 co | Just ty1' <- tcView ty1 = go ty1' ty2 co
+    go ty1 ty2 co | Just ty2' <- tcView ty2 = go ty1 ty2' co
 
-%************************************************************************
-%*                                                                     *
-\section{Different forms of Inst rewrite rules}
-%*                                                                     *
-%************************************************************************
+      -- left-to-right rule with type family head
+    go (TyConApp con args) ty2 co 
+      | isOpenSynTyCon con
+      = mkRewriteFam False con args ty2 co
 
-Splitting of non-terminating given constraints: skolemOccurs
-~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
-This is a preconditioning rule exclusively applied to given constraints.
-Moreover, its rewriting is only temporary, as it is undone by way of
-side-effecting mutable type variables after simplification and constraint
-entailment has been completed.
+      -- right-to-left rule with type family head
+    go ty1 ty2@(TyConApp con args) co 
+      | isOpenSynTyCon con
+      = do { co' <- mkSymEqInstCo co (ty2, ty1)
+           ; mkRewriteFam True con args ty1 co'
+           }
 
-This version is an (attempt at, yet unproven, an) *unflattened* version of
-the SubstL-Ev completion rule.
+      -- no outermost family
+    go ty1 ty2 co
+      = do { (ty1', co1, ty1_eqs, ty1_skolems) <- flattenType inst ty1
+           ; (ty2', co2, ty2_eqs, ty2_skolems) <- flattenType inst ty2
+           ; let ty12_eqs  = ty1_eqs ++ ty2_eqs
+                 sym_co2   = mkSymCoercion co2
+                 eqTys     = (ty1', ty2')
+           ; (co', ty12_eqs') <- adjustCoercions co co1 sym_co2 eqTys ty12_eqs
+           ; eqs <- checkOrientation ty1' ty2' co' inst
+           ; if isLoopyEquality eqs ty12_eqs' 
+             then do { if isWantedCo (tci_co inst)
+                       then
+                          addErrCtxt (ptext (sLit "Rejecting loopy equality")) $
+                            eqInstMisMatch inst
+                       else
+                         warnDroppingLoopyEquality ty1 ty2
+                     ; return ([], emptyVarSet)         -- drop the equality
+                     }
+             else
+               return (eqs ++ ty12_eqs',
+                      ty1_skolems `unionVarSet` ty2_skolems)
+           }
 
-The above rule is essential to catch non-terminating rules that cannot be
-oriented properly, like 
+    mkRewriteFam swapped con args ty2 co
+      = do { (args', cargs, args_eqss, args_skolemss) 
+               <- mapAndUnzip4M (flattenType inst) args
+           ; (ty2', co2, ty2_eqs, ty2_skolems) <- flattenType inst ty2
+           ; let co1       = mkTyConApp con cargs
+                 sym_co2   = mkSymCoercion co2
+                 all_eqs   = concat args_eqss ++ ty2_eqs
+                 eqTys     = (mkTyConApp con args', ty2')
+           ; (co', all_eqs') <- adjustCoercions co co1 sym_co2 eqTys all_eqs
+           ; let thisRewriteFam = RewriteFam 
+                                  { rwi_fam     = con
+                                  , rwi_args    = args'
+                                  , rwi_right   = ty2'
+                                  , rwi_co      = co'
+                                  , rwi_loc     = tci_loc inst
+                                  , rwi_name    = tci_name inst
+                                  , rwi_swapped = swapped
+                                  }
+           ; return $ (thisRewriteFam : all_eqs',
+                       unionVarSets (ty2_skolems:args_skolemss))
+           }
 
-     F a ~ [G (F a)]
- or even
-     a ~ [G a]
+    -- If the original equality has the form a ~ T .. (F ...a...) ..., we will
+    -- have a variable equality with 'a' on the lhs as the first equality.
+    -- Then, check whether 'a' occurs in the lhs of any family equality
+    -- generated by flattening.
+    isLoopyEquality (RewriteVar {rwi_var = tv}:_) eqs
+      = any inRewriteFam eqs
+      where
+        inRewriteFam (RewriteFam {rwi_args = args}) 
+          = tv `elemVarSet` tyVarsOfTypes args
+        inRewriteFam _ = False
+    isLoopyEquality _ _ = False
+
+normDict :: Bool -> Inst -> TcM (Inst, [RewriteInst], TcDictBinds, TyVarSet)
+-- Normalise one dictionary or IP constraint.
+normDict isWanted inst@(Dict {tci_pred = ClassP clas args})
+  = do { (args', cargs, args_eqss, args_skolemss) 
+           <- mapAndUnzip4M (flattenType inst) args
+       ; let rewriteCo = PredTy $ ClassP clas cargs
+             eqs       = concat args_eqss
+             pred'     = ClassP clas args'
+       ; if null eqs
+         then  -- don't generate a binding if there is nothing to flatten
+           return (inst, [], emptyBag, emptyVarSet)
+         else do {
+       ; (inst', bind) <- mkDictBind inst isWanted rewriteCo pred'
+       ; eqs' <- if isWanted then return eqs else mapM wantedToLocal eqs
+       ; return (inst', eqs', bind, unionVarSets args_skolemss)
+       }}
+normDict _isWanted inst
+  = return (inst, [], emptyBag, emptyVarSet)
+-- !!!TODO: Still need to normalise IP constraints.
+
+checkOrientation :: Type -> Type -> EqInstCo -> Inst -> TcM [RewriteInst]
+-- Performs the occurs check, decomposition, and proper orientation
+-- (returns a singleton, or an empty list in case of a trivial equality)
+-- NB: We cannot assume that the two types already have outermost type
+--     synonyms expanded due to the recursion in the case of type applications.
+checkOrientation ty1 ty2 co inst
+  = do { traceTc $ ptext (sLit "checkOrientation of ") <+> 
+                   pprEqInstCo co <+> text "::" <+> 
+                   ppr ty1 <+> text "~" <+> ppr ty2
+       ; eqs <- go ty1 ty2
+       ; traceTc $ ptext (sLit "checkOrientation returns") <+> ppr eqs
+       ; return eqs
+       }
+  where
+      -- look through synonyms
+    go ty1 ty2 | Just ty1' <- tcView ty1 = go ty1' ty2
+    go ty1 ty2 | Just ty2' <- tcView ty2 = go ty1 ty2'
+
+      -- identical types => trivial
+    go ty1 ty2
+      | ty1 `tcEqType` ty2
+      = do { mkIdEqInstCo co ty1
+           ; return []
+           }
 
-The left-to-right orientiation is not suitable because it does not
-terminate. The right-to-left orientation is not suitable because it 
-does not have a type-function on the left. This is undesirable because
-it would hide information. E.g. assume 
+      -- two tvs, left greater => unchanged
+    go ty1@(TyVarTy tv1) ty2@(TyVarTy tv2)
+      | tv1 > tv2
+      = mkRewriteVar False tv1 ty2 co
 
-       instance C [x]
+      -- two tvs, right greater => swap
+      | otherwise
+      = do { co' <- mkSymEqInstCo co (ty2, ty1)
+           ; mkRewriteVar True tv2 ty1 co'
+           }
 
-then rewriting C [G (F a)] to C (F a) is bad because we cannot now
-see that the C [x] instance applies.
+      -- only lhs is a tv => unchanged
+    go ty1@(TyVarTy tv1) ty2
+      | ty1 `tcPartOfType` ty2      -- occurs check!
+      = occurCheckErr ty1 ty2
+      | otherwise 
+      = mkRewriteVar False tv1 ty2 co
+
+      -- only rhs is a tv => swap
+    go ty1 ty2@(TyVarTy tv2)
+      | ty2 `tcPartOfType` ty1      -- occurs check!
+      = occurCheckErr ty2 ty1
+      | otherwise 
+      = do { co' <- mkSymEqInstCo co (ty2, ty1)
+           ; mkRewriteVar True tv2 ty1 co'
+           }
 
-The rule also caters for badly-oriented rules of the form:
+      -- type applications => decompose
+    go ty1 ty2 
+      | Just (ty1_l, ty1_r) <- repSplitAppTy_maybe ty1   -- won't split fam apps
+      , Just (ty2_l, ty2_r) <- repSplitAppTy_maybe ty2
+      = do { (co_l, co_r) <- mkAppEqInstCo co (ty1_l, ty2_l) (ty1_r, ty2_r)
+           ; eqs_l <- checkOrientation ty1_l ty2_l co_l inst
+           ; eqs_r <- checkOrientation ty1_r ty2_r co_r inst
+           ; return $ eqs_l ++ eqs_r
+           }
+-- !!!TODO: would be more efficient to handle the FunApp and the data
+-- constructor application explicitly.
+
+      -- inconsistency => type error
+    go ty1 ty2
+      = ASSERT( (not . isForAllTy $ ty1) && (not . isForAllTy $ ty2) )
+        eqInstMisMatch inst
+
+    mkRewriteVar swapped tv ty co = return [RewriteVar 
+                                            { rwi_var     = tv
+                                            , rwi_right   = ty
+                                            , rwi_co      = co
+                                            , rwi_loc     = tci_loc inst
+                                            , rwi_name    = tci_name inst
+                                            , rwi_swapped = swapped
+                                            }]
+
+flattenType :: Inst     -- context to get location  & name
+            -> Type     -- the type to flatten
+            -> TcM (Type,           -- the flattened type
+                    Coercion,       -- coercion witness of flattening wanteds
+                    [RewriteInst],  -- extra equalities
+                    TyVarSet)       -- new intermediate skolems
+-- Removes all family synonyms from a type by moving them into extra equalities
+flattenType inst ty
+  = go ty
+  where
+      -- look through synonyms
+    go ty | Just ty' <- tcView ty 
+      = do { (ty_flat, co, eqs, skolems) <- go ty'
+           ; if null eqs
+             then     -- unchanged, keep the old type with folded synonyms
+               return (ty, ty, [], emptyVarSet)
+             else 
+               return (ty_flat, co, eqs, skolems)
+           }
 
-     F a ~ G (F a)
+      -- type variable => nothing to do
+    go ty@(TyVarTy _)
+      = return (ty, ty, [] , emptyVarSet)
 
-for which other solutions are possible, but this one will do too.
+      -- type family application 
+      -- => flatten to "gamma :: F t1'..tn' ~ alpha" (alpha & gamma fresh)
+    go ty@(TyConApp con args)
+      | isOpenSynTyCon con
+      = do { (args', cargs, args_eqss, args_skolemss) <- mapAndUnzip4M go args
+           ; alpha <- newFlexiTyVar (typeKind ty)
+           ; let alphaTy = mkTyVarTy alpha
+           ; cotv <- newMetaCoVar (mkTyConApp con args') alphaTy
+           ; let thisRewriteFam = RewriteFam 
+                                  { rwi_fam     = con
+                                  , rwi_args    = args'
+                                  , rwi_right   = alphaTy
+                                  , rwi_co      = mkWantedCo cotv
+                                  , rwi_loc     = tci_loc inst
+                                  , rwi_name    = tci_name inst
+                                  , rwi_swapped = True
+                                  }
+           ; return (alphaTy,
+                     mkTyConApp con cargs `mkTransCoercion` mkTyVarTy cotv,
+                     thisRewriteFam : concat args_eqss,
+                     unionVarSets args_skolemss `extendVarSet` alpha)
+           }           -- adding new unflatten var inst
+
+      -- data constructor application => flatten subtypes
+      -- NB: Special cased for efficiency - could be handled as type application
+    go ty@(TyConApp con args)
+      = do { (args', cargs, args_eqss, args_skolemss) <- mapAndUnzip4M go args
+           ; if null args_eqss
+             then     -- unchanged, keep the old type with folded synonyms
+               return (ty, ty, [], emptyVarSet)
+             else 
+               return (mkTyConApp con args', 
+                       mkTyConApp con cargs,
+                       concat args_eqss,
+                       unionVarSets args_skolemss)
+           }
+
+      -- function type => flatten subtypes
+      -- NB: Special cased for efficiency - could be handled as type application
+    go ty@(FunTy ty_l ty_r)
+      = do { (ty_l', co_l, eqs_l, skolems_l) <- go ty_l
+           ; (ty_r', co_r, eqs_r, skolems_r) <- go ty_r
+           ; if null eqs_l && null eqs_r
+             then     -- unchanged, keep the old type with folded synonyms
+               return (ty, ty, [], emptyVarSet)
+             else 
+               return (mkFunTy ty_l' ty_r', 
+                       mkFunTy co_l co_r,
+                       eqs_l ++ eqs_r, 
+                       skolems_l `unionVarSet` skolems_r)
+           }
 
-It's behavior is:
+      -- type application => flatten subtypes
+    go ty@(AppTy ty_l ty_r)
+      = do { (ty_l', co_l, eqs_l, skolems_l) <- go ty_l
+           ; (ty_r', co_r, eqs_r, skolems_r) <- go ty_r
+           ; if null eqs_l && null eqs_r
+             then     -- unchanged, keep the old type with folded synonyms
+               return (ty, ty, [], emptyVarSet)
+             else 
+               return (mkAppTy ty_l' ty_r', 
+                       mkAppTy co_l co_r, 
+                       eqs_l ++ eqs_r, 
+                       skolems_l `unionVarSet` skolems_r)
+           }
+
+      -- forall type => panic if the body contains a type family
+      -- !!!TODO: As long as the family does not contain a quantified variable
+      --          we might pull it out, but what if it does contain a quantified
+      --          variable???
+    go ty@(ForAllTy _ body)
+      | null (tyFamInsts body)
+      = return (ty, ty, [] , emptyVarSet)
+      | otherwise
+      = panic "TcTyFuns.flattenType: synonym family in a rank-n type"
+
+      -- we should never see a predicate type
+    go (PredTy _)
+      = panic "TcTyFuns.flattenType: unexpected PredType"
+
+adjustCoercions :: EqInstCo            -- coercion of original equality
+                -> Coercion            -- coercion witnessing the left rewrite
+                -> Coercion            -- coercion witnessing the right rewrite
+                -> (Type, Type)        -- types of flattened equality
+                -> [RewriteInst]       -- equalities from flattening
+                -> TcM (EqInstCo,      -- coercion for flattened equality
+                        [RewriteInst]) -- final equalities from flattening
+-- Depending on whether we flattened a local or wanted equality, that equality's
+-- coercion and that of the new equalities produced during flattening are
+-- adjusted .
+adjustCoercions (Left cotv) co1 co2 (ty_l, ty_r) all_eqs
+    -- wanted => generate a fresh coercion variable for the flattened equality
+  = do { cotv' <- newMetaCoVar ty_l ty_r
+       ; writeMetaTyVar cotv $ 
+           (co1 `mkTransCoercion` TyVarTy cotv' `mkTransCoercion` co2)
+       ; return (Left cotv', all_eqs)
+       }
 
-  co : ty1 ~ ty2{F ty1}
-     >-->
-  co         : ty1 ~ ty2{b}
-  sym (F co) : F ty2{b} ~ b
-       where b is a fresh skolem variable
+adjustCoercions co@(Right _) _co1 _co2 _eqTys all_eqs
+    -- local => turn all new equalities into locals and update (but not zonk)
+    --          the skolem
+  = do { all_eqs' <- mapM wantedToLocal all_eqs
+       ; return (co, all_eqs')
+       }
 
-We also cater for the symmetric situation *if* the rule cannot be used as a
-left-to-right rewrite rule.
+mkDictBind :: Inst                 -- original instance
+           -> Bool                 -- is this a wanted contraint?
+           -> Coercion             -- coercion witnessing the rewrite
+           -> PredType             -- coerced predicate
+           -> TcM (Inst,           -- new inst
+                   TcDictBinds)    -- binding for coerced dictionary
+mkDictBind dict isWanted rewriteCo pred
+  = do { dict' <- newDictBndr loc pred
+         -- relate the old inst to the new one
+         -- target_dict = source_dict `cast` st_co
+       ; let (target_dict, source_dict, st_co) 
+               | isWanted  = (dict,  dict', mkSymCoercion rewriteCo)
+               | otherwise = (dict', dict,  rewriteCo)
+                 -- we have
+                 --   co :: dict ~ dict'
+                 -- hence, if isWanted
+                 --      dict  = dict' `cast` sym co
+                 --        else
+                 --      dict' = dict  `cast` co
+             expr      = HsVar $ instToId source_dict
+             cast_expr = HsWrap (WpCast st_co) expr
+             rhs       = L (instLocSpan loc) cast_expr
+             binds     = instToDictBind target_dict rhs
+       ; return (dict', binds)
+       }
+  where
+    loc = tci_loc dict
+
+-- gamma :: Fam args ~ alpha
+-- => alpha :: Fam args ~ alpha, with alpha := Fam args
+--    (the update of alpha will not be apparent during propagation, as we
+--    never follow the indirections of meta variables; it will be revealed
+--    when the equality is zonked)
+wantedToLocal :: RewriteInst -> TcM RewriteInst
+wantedToLocal eq@(RewriteFam {rwi_fam   = fam, 
+                              rwi_args  = args, 
+                              rwi_right = alphaTy@(TyVarTy alpha)})
+  = do { writeMetaTyVar alpha (mkTyConApp fam args)
+       ; return $ eq {rwi_co = mkGivenCo alphaTy}
+       }
+wantedToLocal _ = panic "TcTyFuns.wantedToLocal"
+\end{code}
 
-We also return an action (b := ty1) which is used to eliminate b 
-after the dust of normalisation with the completed rewrite system
-has settled.
 
-A subtle point of this transformation is that both coercions in the results
-are strictly speaking incorrect.  However, they are correct again after the
-action {B := ty1} has removed the skolem again.  This happens immediately
-after constraint entailment has been checked; ie, code outside of the
-simplification and entailment checking framework will never see these
-temporarily incorrect coercions.
+%************************************************************************
+%*                                                                     *
+               Propagation of equalities
+%*                                                                     *
+%************************************************************************
 
-NB: We perform this transformation for multiple occurences of ty1 under one
-    or multiple family applications on the left-hand side at once (ie, the
-    rule doesn't need to be applied multiple times at a single inst).  As a 
-    result we can get two or more insts back.
+Apply the propagation rules exhaustively.
 
 \begin{code}
-skolemOccurs :: PrecondRule
-skolemOccurs insts
-  = do { (instss, undoSkolems) <- mapAndUnzipM oneSkolemOccurs insts
-       ; return (concat instss, sequence_ undoSkolems)
+propagate :: [RewriteInst] -> EqConfig -> TcM EqConfig
+propagate []       eqCfg = return eqCfg
+propagate (eq:eqs) eqCfg
+  = do { optEqs <- applyTop eq
+       ; case optEqs of
+
+              -- Top applied to 'eq' => retry with new equalities
+           Just (eqs2, skolems2) 
+             -> propagate (eqs2 ++ eqs) (eqCfg `addSkolems` skolems2)
+
+              -- Top doesn't apply => try subst rules with all other
+              --   equalities, after that 'eq' can go into the residual list
+           Nothing
+             -> do { (eqs', eqCfg') <- applySubstRules eq eqs eqCfg
+                   ; propagate eqs' (eqCfg' `addEq` eq)
+                   }
+   }
+
+applySubstRules :: RewriteInst                    -- currently considered eq
+                -> [RewriteInst]                  -- todo eqs list
+                -> EqConfig                       -- residual
+                -> TcM ([RewriteInst], EqConfig)  -- new todo & residual
+applySubstRules eq todoEqs (eqConfig@EqConfig {eqs = resEqs})
+  = do { (newEqs_t, unchangedEqs_t, skolems_t) <- mapSubstRules eq todoEqs
+       ; (newEqs_r, unchangedEqs_r, skolems_r) <- mapSubstRules eq resEqs
+       ; return (newEqs_t ++ newEqs_r ++ unchangedEqs_t,
+                 eqConfig {eqs = unchangedEqs_r} 
+                   `addSkolems` (skolems_t `unionVarSet` skolems_r))
        }
-  where
-    oneSkolemOccurs inst
-      = ASSERT( isEqInst inst )
-        isRewriteRule (eqInstLeftTy inst) (eqInstRightTy inst)
-      where
-
-        -- look through synonyms
-        isRewriteRule ty1 ty2 | Just ty1' <- tcView ty1 = isRewriteRule ty1' ty2
-        isRewriteRule ty1 ty2 | Just ty2' <- tcView ty2 = isRewriteRule ty1 ty2'
 
-        -- left-to-right rule with type family head
-        isRewriteRule ty1@(TyConApp con _) ty2
-          | isOpenSynTyCon con
-          = breakRecursion ty1 ty2 False    -- not swapped
+mapSubstRules :: RewriteInst     -- try substituting this equality
+              -> [RewriteInst]   -- into these equalities
+              -> TcM ([RewriteInst], [RewriteInst], TyVarSet)
+mapSubstRules eq eqs
+  = do { (newEqss, unchangedEqss, skolemss) <- mapAndUnzip3M (substRules eq) eqs
+       ; return (concat newEqss, concat unchangedEqss, unionVarSets skolemss)
+       }
+  where
+    substRules eq1 eq2
+      = do {   -- try the SubstFam rule
+             optEqs <- applySubstFam eq1 eq2
+           ; case optEqs of
+               Just (eqs, skolems) -> return (eqs, [], skolems)
+               Nothing             -> do 
+           {   -- try the SubstVarVar rule
+             optEqs <- applySubstVarVar eq1 eq2
+           ; case optEqs of
+               Just (eqs, skolems) -> return (eqs, [], skolems)
+               Nothing             -> do 
+           {   -- try the SubstVarFam rule
+             optEqs <- applySubstVarFam eq1 eq2
+           ; case optEqs of
+               Just eq -> return ([eq], [], emptyVarSet)
+               Nothing -> return ([], [eq2], emptyVarSet)
+                 -- if no rule matches, we return the equlity we tried to
+                 -- substitute into unchanged
+           }}}
+\end{code}
 
-        -- left-to-right rule with type variable head
-        isRewriteRule ty1@(TyVarTy _) ty2
-          = breakRecursion ty1 ty2 False    -- not swapped
+Attempt to apply the Top rule.  The rule is
 
-        -- right-to-left rule with type family head
-        isRewriteRule ty1 ty2@(TyConApp con _)
-          | isOpenSynTyCon con
-          = breakRecursion ty2 ty1 True     -- swapped
+  co :: F t1..tn ~ t
+  =(Top)=>
+  co' :: [s1/x1, .., sm/xm]s ~ t with co = g s1..sm |> co'  
 
-        -- right-to-left rule with type variable head
-        isRewriteRule ty1 ty2@(TyVarTy _)
-          = breakRecursion ty2 ty1 True     -- swapped
+where g :: forall x1..xm. F u1..um ~ s and [s1/x1, .., sm/xm]u1 == t1.
 
-        -- this equality is not a rewrite rule => ignore
-        isRewriteRule _ _ = return ([inst], return ())
+Returns Nothing if the rule could not be applied.  Otherwise, the resulting
+equality is normalised and a list of the normal equalities is returned.
 
+\begin{code}
+applyTop :: RewriteInst -> TcM (Maybe ([RewriteInst], TyVarSet))
 
-        breakRecursion pat body swapped
-          | null tysToLiftOut 
-          = return ([inst], return ())
-          | otherwise
-          = do { skTvs <- mapM (newMetaTyVar TauTv . typeKind) tysToLiftOut
-               ; let skTvs_tysTLO   = zip skTvs tysToLiftOut
-                     insertSkolems = return . replace skTvs_tysTLO
-               ; (_, body') <- tcGenericNormaliseFamInst insertSkolems body
-               ; inst' <- if swapped then mkEqInst (EqPred body' pat) co
-                                     else mkEqInst (EqPred pat body') co
-                                     -- ensure to reconstruct the inst in the
-                                     -- original orientation
-               ; (insts, undoSk) <- mapAndUnzipM (mkSkolemInst inst') 
-                                                 skTvs_tysTLO
-               ; return (inst':insts, sequence_ undoSk)
-               }
-          where
-            co  = eqInstCoercion inst
-
-            -- all subtypes that are (1) type family instances and (2) contain
-            -- the lhs type as part of the type arguments of the type family
-            -- constructor 
-            tysToLiftOut = [mkTyConApp tc tys | (tc, tys) <- tyFamInsts body
-                                              , any (pat `tcPartOfType`) tys]
-
-            replace :: [(TcTyVar, Type)] -> Type -> Maybe (Type, Coercion)
-            replace []                   _  = Nothing
-            replace ((skTv, tyTLO):rest) ty 
-              | tyTLO `tcEqType` ty         = Just (mkTyVarTy skTv, undefined)
-              | otherwise                   = replace rest ty
-
-            -- create the EqInst for the equality determining the skolem and a
-            -- TcM action undoing the skolem introduction
-            mkSkolemInst inst' (skTv, tyTLO)
-              = do { (co, tyLiftedOut) <- tcEqInstNormaliseFamInst inst' tyTLO
-                   ; inst <- mkEqInst (EqPred tyLiftedOut (mkTyVarTy skTv)) 
-                                      (mkGivenCo $ mkSymCoercion (fromACo co))
-                   ; return (inst, writeMetaTyVar skTv tyTLO) 
+applyTop eq@(RewriteFam {rwi_fam = fam, rwi_args = args})
+  = do { optTyCo <- tcUnfoldSynFamInst (TyConApp fam args)
+       ; case optTyCo of
+           Nothing                -> return Nothing
+           Just (lhs, rewrite_co) 
+             -> do { co' <- mkRightTransEqInstCo co rewrite_co (lhs, rhs)
+                   ; eq' <- deriveEqInst eq lhs rhs co'
+                   ; liftM Just $ normEqInst eq'
                    }
+       }
+  where
+    co  = rwi_co eq
+    rhs = rwi_right eq
+
+applyTop _ = return Nothing
 \end{code}
 
+Attempt to apply the SubstFam rule.  The rule is
 
-Removal of trivial equalities: trivialRule
-~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
-The following rules exploits the reflexivity of equality:
+  co1 :: F t1..tn ~ t  &  co2 :: F t1..tn ~ s
+  =(SubstFam)=>
+  co1 :: F t1..tn ~ t  &  co2' :: t ~ s with co2 = co1 |> co2'
 
-  (Trivial)
-    g1 : t ~ t
-      >-->
-    g1 := t
+where co1 may be a wanted only if co2 is a wanted, too.
+
+Returns Nothing if the rule could not be applied.  Otherwise, the equality
+co2' is normalised and a list of the normal equalities is returned.  (The
+equality co1 is not returned as it remain unaltered.)
 
 \begin{code}
-trivialRule :: IdemRewriteRule
-trivialRule insts 
-  = liftM catMaybes $ mappM trivial insts
+applySubstFam :: RewriteInst 
+              -> RewriteInst 
+              -> TcM (Maybe ([RewriteInst], TyVarSet))
+applySubstFam eq1@(RewriteFam {rwi_fam = fam1, rwi_args = args1})
+              eq2@(RewriteFam {rwi_fam = fam2, rwi_args = args2})
+  | fam1 == fam2 && tcEqTypes args1 args2 &&
+    (isWantedRewriteInst eq2 || not (isWantedRewriteInst eq1))
+-- !!!TODO: tcEqTypes is insufficient as it does not look through type synonyms
+-- !!!Check whether anything breaks by making tcEqTypes look through synonyms.
+-- !!!Should be ok and we don't want three type equalities.
+  = do { co2' <- mkRightTransEqInstCo co2 co1 (lhs, rhs)
+       ; eq2' <- deriveEqInst eq2 lhs rhs co2'
+       ; liftM Just $ normEqInst eq2'
+       }
   where
-    trivial inst
-      | ASSERT( isEqInst inst )
-        ty1 `tcEqType` ty2
-      = do { eitherEqInst inst
-              (\cotv -> writeMetaTyVar cotv ty1) 
-              (\_    -> return ())
-          ; return Nothing
-          }
-      | otherwise
-      = return $ Just inst
-      where
-        ty1 = eqInstLeftTy  inst
-        ty2 = eqInstRightTy inst
+    lhs = rwi_right eq1
+    rhs = rwi_right eq2
+    co1 = eqInstCoType (rwi_co eq1)
+    co2 = rwi_co eq2
+applySubstFam _ _ = return Nothing
 \end{code}
 
+Attempt to apply the SubstVarVar rule.  The rule is
 
-Decomposition of data type constructors: decompRule
-~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
-Whenever, the same *data* constructors occurs on both sides of an equality, we
-can decompose as in standard unification.
-
-  (Decomp)
-    g1 : T cs ~ T ds
-      >-->
-    g21 : c1 ~ d1, ..., g2n : cn ~ dn
-    g1 := T g2s
+  co1 :: x ~ t  &  co2 :: x ~ s
+  =(SubstVarVar)=>
+  co1 :: x ~ t  &  co2' :: t ~ s with co2 = co1 |> co2'
 
-Works also for the case where T is actually an application of a type family
-constructor to a set of types, provided the applications on both sides of the
-~ are identical; see also Note [OpenSynTyCon app] in TcUnify.
+where co1 may be a wanted only if co2 is a wanted, too.
 
-We guarantee to raise an error for any inconsistent equalities; 
-cf Note [Inconsistencies in equality constraints].
+Returns Nothing if the rule could not be applied.  Otherwise, the equality
+co2' is normalised and a list of the normal equalities is returned.  (The
+equality co1 is not returned as it remain unaltered.)
 
 \begin{code}
-decompRule :: RewriteRule
-decompRule insts 
-  = do { (insts, changed) <- mapAndUnzipM decomp insts
-       ; return (concat insts, or changed)
+applySubstVarVar :: RewriteInst 
+                 -> RewriteInst 
+                 -> TcM (Maybe ([RewriteInst], TyVarSet))
+applySubstVarVar eq1@(RewriteVar {rwi_var = tv1})
+                 eq2@(RewriteVar {rwi_var = tv2})
+  | tv1 == tv2 &&
+    (isWantedRewriteInst eq2 || not (isWantedRewriteInst eq1))
+  = do { co2' <- mkRightTransEqInstCo co2 co1 (lhs, rhs)
+       ; eq2' <- deriveEqInst eq2 lhs rhs co2'
+       ; liftM Just $ normEqInst eq2'
        }
   where
-    decomp inst
-      = ASSERT( isEqInst inst )
-        go (eqInstLeftTy inst) (eqInstRightTy inst)
-      where
-        go ty1 ty2             
-          | Just ty1' <- tcView ty1 = go ty1' ty2 
-          | Just ty2' <- tcView ty2 = go ty1 ty2' 
-
-        go (TyConApp con1 tys1) (TyConApp con2 tys2)
-          | con1 == con2 && identicalHead
-          = mkArgInsts (mkTyConApp con1) tys1 tys2
-
-          | con1 /= con2 && not (isOpenSynTyCon con1 || isOpenSynTyCon con2)
-            -- not matching data constructors (of any flavour) are bad news
-          = eqInstMisMatch inst
-          where
-            n             = tyConArity con1
-            (idxTys1, _)  = splitAt n tys1
-            (idxTys2, _)  = splitAt n tys2
-            identicalHead = not (isOpenSynTyCon con1) ||
-                            idxTys1 `tcEqTypes` idxTys2
-
-        go (FunTy fun1 arg1) (FunTy fun2 arg2)
-          = mkArgInsts (\[funCo, argCo] -> mkFunTy funCo argCo) [fun1, arg1]
-                                                                [fun2, arg2]
-
-       -- Applications need a bit of care!
-       -- They can match FunTy and TyConApp, so use splitAppTy_maybe
-        go (AppTy s1 t1) ty2
-          | Just (s2, t2) <- tcSplitAppTy_maybe ty2
-          = mkArgInsts (\[s, t] -> mkAppTy s t) [s1, t1] [s2, t2]
-
-        -- Symmetric case
-        go ty1 (AppTy s2 t2)
-          | Just (s1, t1) <- tcSplitAppTy_maybe ty1
-          = mkArgInsts (\[s, t] -> mkAppTy s t) [s1, t1] [s2, t2]
-
-        -- We already covered all the consistent cases of rigid types on both
-        -- sides; so, if we see two rigid types here, we discovered an
-        -- inconsistency. 
-        go ty1 ty2 
-          | isRigid ty1 && isRigid ty2
-          = eqInstMisMatch inst
-
-        -- We can neither assert consistency nor inconsistency => defer
-        go _ _ = return ([inst], False)
-
-        isRigid (TyConApp con _) = not (isOpenSynTyCon con)
-        isRigid (FunTy _ _)      = True
-        isRigid (AppTy _ _)      = True
-        isRigid _                = False
-
-        -- Create insts for matching argument positions (ie, the bit after
-        -- '>-->' in the rule description above)
-        mkArgInsts con tys1 tys2
-          = do { cos <- eitherEqInst inst
-                          -- old_co := Con1 cos
-                          (\old_covar ->
-                            do { cotvs <- zipWithM newMetaCoVar tys1 tys2
-                               ; let cos = map mkTyVarTy cotvs
-                               ; writeMetaTyVar old_covar (con cos)
-                               ; return $ map mkWantedCo cotvs
-                               })
-                          -- co_i := Con_i old_co
-                          (\old_co -> 
-                            return $ map mkGivenCo $
-                                         mkRightCoercions (length tys1) old_co)
-               ; insts <- zipWithM mkEqInst (zipWith EqPred tys1 tys2) cos
-               ; traceTc (text "decomp identicalHead" <+> ppr insts) 
-               ; return (insts, not $ null insts) 
-               }
+    lhs = rwi_right eq1
+    rhs = rwi_right eq2
+    co1 = eqInstCoType (rwi_co eq1)
+    co2 = rwi_co eq2
+applySubstVarVar _ _ = return Nothing
 \end{code}
 
+Attempt to apply the SubstVarFam rule.  The rule is
+
+  co1 :: x ~ t  &  co2 :: F s1..sn ~ s
+  =(SubstVarFam)=>
+  co1 :: x ~ t  &  co2' :: [t/x](F s1..sn) ~ s 
+    with co2 = [co1/x](F s1..sn) |> co2'
 
-Rewriting with type instances: topRule
-~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
-We use (toplevel) type instances to normalise both sides of equalities.
+where x occurs in F s1..sn. (co1 may be local or wanted.)
 
-  (Top)
-    g1 : t ~ s
-      >--> co1 :: t ~ t' / co2 :: s ~ s'
-    g2 : t' ~ s'
-    g1 := co1 * g2 * sym co2
+Returns Nothing if the rule could not be applied.  Otherwise, the equality
+co2' is returned.  (The equality co1 is not returned as it remain unaltered.)
 
 \begin{code}
-topRule :: RewriteRule
-topRule insts 
-  =  do { (insts, changed) <- mapAndUnzipM top insts
-       ; return (insts, or changed)
-       }
+applySubstVarFam :: RewriteInst -> RewriteInst -> TcM (Maybe RewriteInst)
+applySubstVarFam eq1@(RewriteVar {rwi_var = tv1})
+                 eq2@(RewriteFam {rwi_fam = fam2, rwi_args = args2})
+  | tv1 `elemVarSet` tyVarsOfTypes args2
+  = do { let co1Subst = substTyWith [tv1] [co1] (mkTyConApp fam2 args2)
+             args2'   = substTysWith [tv1] [rhs1] args2
+             lhs2     = mkTyConApp fam2 args2'
+       ; co2' <- mkRightTransEqInstCo co2 co1Subst (lhs2, rhs2)
+       ; return $ Just (eq2 {rwi_args = args2', rwi_co = co2'})
+       }
   where
-    top inst
-      = ASSERT( isEqInst inst )
-        do { (coi1, ty1') <- tcNormaliseFamInst ty1
-          ; (coi2, ty2') <- tcNormaliseFamInst ty2
-          ; case (coi1, coi2) of
-              (IdCo, IdCo) -> return (inst, False)
-              _            -> 
-                 do { wg_co <- 
-                       eitherEqInst inst
-                         -- old_co = co1 * new_co * sym co2
-                         (\old_covar ->
-                           do { new_cotv <- newMetaCoVar ty1' ty2'
-                              ; let new_co  = mkTyVarTy new_cotv
-                                    old_coi = coi1 `mkTransCoI` 
-                                              ACo new_co `mkTransCoI` 
-                                              (mkSymCoI coi2)
-                              ; writeMetaTyVar old_covar (fromACo old_coi)
-                              ; return $ mkWantedCo new_cotv
-                              })
-                         -- new_co = sym co1 * old_co * co2
-                         (\old_co -> 
-                           return $ 
-                             mkGivenCo $ 
-                               fromACo $ 
-                                 mkSymCoI coi1 `mkTransCoI` 
-                                 ACo old_co `mkTransCoI` coi2) 
-                   ; new_inst <- mkEqInst (EqPred ty1' ty2') wg_co 
-                   ; return (new_inst, True)
-                   }
-            }
-      where
-        ty1 = eqInstLeftTy  inst
-       ty2 = eqInstRightTy inst
+    rhs1 = rwi_right eq1
+    rhs2 = rwi_right eq2
+    co1  = eqInstCoType (rwi_co eq1)
+    co2  = rwi_co eq2
+applySubstVarFam _ _ = return Nothing
 \end{code}
 
 
-Rewriting with equalities: substRule
-~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
-From a set of insts, use all insts that can be read as rewrite rules to
-rewrite the types in all other insts.
-
-  (Subst)
-    g : F c ~ t,
-    forall g1 : s1{F c} ~ s2{F c}
-      >-->
-    g2 : s1{t} ~ s2{t}
-    g1 := s1{g} * g2  * sym s2{g}  <=>  g2 := sym s1{g} * g1 * s2{g}
+%************************************************************************
+%*                                                                     *
+               Finalisation of equalities
+%*                                                                     *
+%************************************************************************
 
-Alternatively, the rewrite rule may have the form (g : a ~ t).
+Exhaustive substitution of all variable equalities of the form co :: x ~ t
+(both local and wanted) into the left-hand sides of all other equalities.  This
+may lead to recursive equalities; i.e., (1) we need to apply the substitution
+implied by one variable equality exhaustively before turning to the next and
+(2) we need an occurs check.
 
-To avoid having to swap rules of the form (g : t ~ F c) and (g : t ~ a),
-where t is neither a variable nor a type family application, we use them for
-rewriting from right-to-left.  However, it is crucial to only apply rules
-from right-to-left if they cannot be used left-to-right.
+We also apply the same substitutions to the local and wanted class and IP
+dictionaries.
 
-The workhorse is substInst, which performs an occurs check before actually
-using an equality for rewriting.  If the type pattern occurs in the type we
-substitute for the pattern, normalisation would diverge.
+NB: Given that we apply the substitution corresponding to a single equality
+exhaustively, before turning to the next, and because we eliminate recursive
+equalities, all opportunities for subtitution will have been exhausted after
+we have considered each equality once.
 
 \begin{code}
-substRule :: RewriteRule
-substRule insts = tryAllInsts insts []
+substitute :: [RewriteInst]       -- equalities
+           -> [Inst]              -- local class dictionaries
+           -> [Inst]              -- wanted class dictionaries
+           -> TcM ([RewriteInst], -- equalities after substitution
+                   TcDictBinds,   -- all newly generated dictionary bindings
+                   [Inst],        -- local dictionaries after substitution
+                   [Inst])        -- wanted dictionaries after substitution
+substitute eqs locals wanteds = subst eqs [] emptyBag locals wanteds
   where
-    -- for every inst check whether it can be used to rewrite the others
-    -- (we make an effort to keep the insts in order; it makes debugging
-    -- easier)
-    tryAllInsts []           triedInsts = return (reverse triedInsts, False)
-    tryAllInsts (inst:insts) triedInsts
-      = do { (insts', changed) <- substInst inst (reverse triedInsts ++ insts)
-           ; if changed then return (insertAt (length triedInsts) inst insts',
-                                     True)
-                        else tryAllInsts insts (inst:triedInsts)
+    subst [] res binds locals wanteds 
+      = return (res, binds, locals, wanteds)
+    subst (eq@(RewriteVar {rwi_var = tv, rwi_right = ty, rwi_co = co}):eqs) 
+          res binds locals wanteds
+      = do { traceTc $ ptext (sLit "TcTyFuns.substitute:") <+> ppr eq
+           ; let coSubst = zipOpenTvSubst [tv] [eqInstCoType co]
+                 tySubst = zipOpenTvSubst [tv] [ty]
+           ; eqs'               <- mapM (substEq eq coSubst tySubst) eqs
+           ; res'               <- mapM (substEq eq coSubst tySubst) res
+           ; (lbinds, locals')  <- mapAndUnzipM 
+                                     (substDict eq coSubst tySubst False) 
+                                     locals
+           ; (wbinds, wanteds') <- mapAndUnzipM 
+                                     (substDict eq coSubst tySubst True) 
+                                     wanteds
+           ; let binds' = unionManyBags $ binds : lbinds ++ wbinds
+           ; subst eqs' (eq:res') binds' locals' wanteds'
+           }
+    subst (eq:eqs) res binds locals wanteds
+      = subst eqs (eq:res) binds locals wanteds
+
+      -- We have, co :: tv ~ ty 
+      -- => apply [ty/tv] to right-hand side of eq2
+      --    (but only if tv actually occurs in the right-hand side of eq2)
+    substEq (RewriteVar {rwi_var = tv, rwi_right = ty}) 
+            coSubst tySubst eq2
+      | tv `elemVarSet` tyVarsOfType (rwi_right eq2)
+      = do { let co1Subst = mkSymCoercion $ substTy coSubst (rwi_right eq2)
+                 right2'  = substTy tySubst (rwi_right eq2)
+                 left2    = case eq2 of
+                              RewriteVar {rwi_var = tv2}   -> mkTyVarTy tv2
+                              RewriteFam {rwi_fam = fam,
+                                          rwi_args = args} ->mkTyConApp fam args
+           ; co2' <- mkLeftTransEqInstCo (rwi_co eq2) co1Subst (left2, right2')
+           ; case eq2 of
+               RewriteVar {rwi_var = tv2} | tv2 `elemVarSet` tyVarsOfType ty
+                 -> occurCheckErr left2 right2'
+               _ -> return $ eq2 {rwi_right = right2', rwi_co = co2'}
            }
-      where
-        insertAt n x xs = let (xs1, xs2) = splitAt n xs
-                          in xs1 ++ [x] ++ xs2
-
--- Use the given inst as a rewrite rule to normalise the insts in the second
--- argument.  Don't do anything if the inst cannot be used as a rewrite rule,
--- but do apply it right-to-left, if possible, and if it cannot be used
--- left-to-right. 
---
-substInst :: Inst -> [Inst] -> TcM ([Inst], Bool)
-substInst inst insts
-  = ASSERT( isEqInst inst )
-    go (eqInstLeftTy inst) (eqInstRightTy inst) (eqInstType inst)
-  where
-    -- look through synonyms
-    go ty1 ty2 co | Just ty1' <- tcView ty1 = go ty1' ty2 co
-    go ty1 ty2 co | Just ty2' <- tcView ty2 = go ty1 ty2' co
-
-    -- rewrite type family applications
-    go ty1@(TyConApp con _) ty2 co 
-      | isOpenSynTyCon con
-      = substEquality (ty1, ty2, co) insts
-
-    -- rewrite skolems
-    go ty1@(TyVarTy tv) ty2 co 
-      | isSkolemTyVar tv
-      = substEquality (ty1, ty2, co) insts
 
-    -- rewrite type family applications from right-to-left, only after
-    -- having checked whether we can work left-to-right
-    go ty1 ty2@(TyConApp con _) co 
-      | isOpenSynTyCon con
-      = substEquality (ty2, ty1, mkSymCoercion co) insts
-
-    -- rewrite skolems from right-to-left, only after having checked
-    -- whether we can work left-to-right 
-    go ty1 ty2@(TyVarTy tv) co 
-      | isSkolemTyVar tv
-      = substEquality (ty2, ty1, mkSymCoercion co) insts
-
-    go _ _ _ = return (insts, False)
-
-    substEquality :: (TcType,   -- rewrite rule lhs
-                      TcType,   -- rewrite rule rhs
-                      TcType)   -- coercion witnessing the rewrite rule
-                  -> [Inst]     -- insts to rewrite
-                  -> TcM ([Inst], Bool)
-    substEquality eqRule@(pat, rhs, _) insts
-      | pat `tcPartOfType` rhs      -- occurs check!
-      = occurCheckErr pat rhs
-      | otherwise
-      = do { (insts', changed) <- mapAndUnzipM substOne insts
-           ; return (insts', or changed)
+      -- unchanged
+    substEq _ _ _ eq2
+      = return eq2
+
+      -- We have, co :: tv ~ ty 
+      -- => apply [ty/tv] to dictionary predicate
+      --    (but only if tv actually occurs in the predicate)
+    substDict (RewriteVar {rwi_var = tv}) 
+              coSubst tySubst isWanted dict
+      | isClassDict dict
+      , tv `elemVarSet` tyVarsOfPred (tci_pred dict)
+      = do { let co1Subst = PredTy (substPred coSubst (tci_pred dict))
+                 pred'    = substPred tySubst (tci_pred dict)
+           ; (dict', binds) <- mkDictBind dict isWanted co1Subst pred'
+           ; return (binds, dict')
            }
-      where
-        substOne inst
-          = ASSERT( isEqInst inst )
-            do { (coi1, ty1') <- tcEqRuleNormaliseFamInst eqRule ty1
-              ; (coi2, ty2') <- tcEqRuleNormaliseFamInst eqRule ty2
-              ; case (coi1, coi2) of
-               (IdCo, IdCo) -> return (inst, False)
-               _            -> 
-                 do { gw_co <- 
-                         eitherEqInst inst
-                          -- old_co := co1 * new_co * sym co2
-                          (\old_covar ->
-                            do { new_cotv <- newMetaCoVar ty1' ty2'
-                               ; let new_co  = mkTyVarTy new_cotv
-                                     old_coi = coi1 `mkTransCoI` 
-                                                ACo new_co `mkTransCoI` 
-                                                (mkSymCoI coi2)
-                               ; writeMetaTyVar old_covar (fromACo old_coi)
-                               ; return $ mkWantedCo new_cotv
-                               })
-                          -- new_co := sym co1 * old_co * co2
-                          (\old_co -> 
-                             return $ 
-                               mkGivenCo $ 
-                                 fromACo $ 
-                                   mkSymCoI coi1 `mkTransCoI` 
-                                   ACo old_co `mkTransCoI` coi2)
-                    ; new_inst <- mkEqInst (EqPred ty1' ty2') gw_co
-                    ; return (new_inst, True)
-                    }
-              }
-         where 
-            ty1 = eqInstLeftTy  inst
-            ty2 = eqInstRightTy inst
+
+      -- unchanged
+    substDict _ _ _ _ dict
+      = return (emptyBag, dict)
+-- !!!TODO: Still need to substitute into IP constraints.
 \end{code}
 
+For any *wanted* variable equality of the form co :: alpha ~ t or co :: a ~
+alpha, we instantiate alpha with t or a, respectively, and set co := id.
+Return all remaining wanted equalities.  The Boolean result component is True
+if at least one instantiation of a flexible was performed.
 
-Instantiate meta variables: unifyMetaRule
-~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~
-If an equality equates a meta type variable with a type, we simply instantiate
-the meta variable.
+\begin{code}
+instantiateAndExtract :: [RewriteInst] -> TcM ([Inst], Bool)
+instantiateAndExtract eqs
+  = do { let wanteds = filter (isWantedCo . rwi_co) eqs
+       ; wanteds' <- mapM inst wanteds
+       ; let residuals = catMaybes wanteds'
+             improved  = length wanteds /= length residuals
+       ; residuals' <- mapM rewriteInstToInst residuals
+       ; return (residuals', improved)
+       }
+  where
+    inst eq@(RewriteVar {rwi_var = tv1, rwi_right = ty2, rwi_co = co})
 
-  (UnifyMeta)
-    g : alpha ~ t
-      >-->
-    alpha := t
-    g     := t
+        -- co :: alpha ~ t
+      | isMetaTyVar tv1
+      = doInst (rwi_swapped eq) tv1 ty2 co eq
 
-Meta variables can only appear in wanted constraints, and this rule should
-only be applied to wanted constraints.  We also know that t definitely is
-distinct from alpha (as the trivialRule) has been run on the insts beforehand.
+        -- co :: a ~ alpha
+      | Just tv2 <- tcGetTyVar_maybe ty2
+      , isMetaTyVar tv2
+      = doInst (not $ rwi_swapped eq) tv2 (mkTyVarTy tv1) co eq
 
-NB: We cannot assume that meta tyvars are empty.  They may have been updated
-by another inst in the currently processed wanted list.  We need to be very
-careful when updateing type variables (see TcUnify.uUnfilledVar), but at least
-we know that we have no boxes.  It's unclear that it would be an advantage to
-common up the code in TcUnify and the code below.  Firstly, we don't want
-calls to TcUnify.defer_unification here, and secondly, TcUnify import the
-current module, so we would have to move everything here (Yuk!) or to
-TcMType.  Besides, the code here is much simpler due to the lack of boxes.
+    inst eq = return $ Just eq
 
-\begin{code}
-unifyMetaRule :: RewriteRule
-unifyMetaRule insts 
-  = do { (insts', changed) <- mapAndUnzipM unifyMeta insts
-       ; return (concat insts', or changed)
-       }
-  where
-    unifyMeta inst
-      = ASSERT( isEqInst inst )
-        go (eqInstLeftTy inst) (eqInstRightTy inst)
-           (fromWantedCo "unifyMetaRule" $ eqInstCoercion inst)
+    doInst _swapped _tv _ty (Right ty) _eq 
+      = pprPanic "TcTyFuns.doInst: local eq: " (ppr ty)
+    doInst swapped tv ty (Left cotv) eq
+      = do { lookupTV <- lookupTcTyVar tv
+           ; uMeta swapped tv lookupTV ty cotv
+           }
       where
-        go ty1 ty2 cotv
-          | Just ty1' <- tcView ty1 = go ty1' ty2 cotv
-          | Just ty2' <- tcView ty2 = go ty1 ty2' cotv
-
-          | TyVarTy tv1 <- ty1
-          , isMetaTyVar tv1         = do { lookupTV <- lookupTcTyVar tv1
-                                         ; uMeta False tv1 lookupTV ty2 cotv
-                                         }
-          | TyVarTy tv2 <- ty2
-          , isMetaTyVar tv2         = do { lookupTV <- lookupTcTyVar tv2
-                                         ; uMeta True tv2 lookupTV ty1 cotv
-                                         }
-          | otherwise               = return ([inst], False) 
-
         -- meta variable has been filled already
-        -- => ignore this inst (we'll come around again, after zonking)
+        -- => ignore (must be a skolem that was introduced by flattening locals)
         uMeta _swapped _tv (IndirectTv _) _ty _cotv
-          = return ([inst], False)
-
-        -- signature skolem meets non-variable type
-        -- => cannot update!
-        uMeta _swapped _tv (DoneTv (MetaTv (SigTv _) _)) ty _cotv
-          | not $ isTyVarTy ty
-          = return ([inst], False)
+          = return Nothing
 
         -- type variable meets type variable
         -- => check that tv2 hasn't been updated yet and choose which to update
-       uMeta swapped tv1 (DoneTv details1) (TyVarTy tv2) cotv
-         = do { lookupTV2 <- lookupTcTyVar tv2
+        uMeta swapped tv1 (DoneTv details1) (TyVarTy tv2) cotv
+          | tv1 == tv2
+          = panic "TcTyFuns.uMeta: normalisation shouldn't allow x ~ x"
+
+          | otherwise
+          = do { lookupTV2 <- lookupTcTyVar tv2
                ; case lookupTV2 of
-                   IndirectTv ty  -> uMeta swapped tv1 (DoneTv details1) ty cotv
+                   IndirectTv ty   -> 
+                     uMeta swapped tv1 (DoneTv details1) ty cotv
                    DoneTv details2 -> 
                      uMetaVar swapped tv1 details1 tv2 details2 cotv
-              }
+               }
+
+        ------ Beyond this point we know that ty2 is not a type variable
+
+        -- signature skolem meets non-variable type
+        -- => cannot update (retain the equality)!
+        uMeta _swapped _tv (DoneTv (MetaTv (SigTv _) _)) _non_tv_ty _cotv
+          = return $ Just eq
 
         -- updatable meta variable meets non-variable type
         -- => occurs check, monotype check, and kinds match check, then update
-       uMeta swapped tv (DoneTv (MetaTv _ ref)) ty cotv
-         = do { mb_ty' <- checkTauTvUpdate tv ty    -- occurs + monotype check
+        uMeta swapped tv (DoneTv (MetaTv _ ref)) non_tv_ty cotv
+          = do {   -- occurs + monotype check
+               ; mb_ty' <- checkTauTvUpdate tv non_tv_ty    
+                             
                ; case mb_ty' of
-                   Nothing  -> return ([inst], False)  -- tv occurs in faminst
+                   Nothing  -> 
+                     -- normalisation shouldn't leave families in non_tv_ty
+                     panic "TcTyFuns.uMeta: unexpected synonym family"
                    Just ty' ->
                      do { checkUpdateMeta swapped tv ref ty'  -- update meta var
-                       ; writeMetaTyVar cotv ty'             -- update co var
-                       ; return ([], True)
+                        ; writeMetaTyVar cotv ty'             -- update co var
+                        ; return Nothing
                         }
-              }
+               }
 
-        uMeta _ _ _ _ _ = panic "uMeta"
+        uMeta _ _ _ _ _ = panic "TcTyFuns.uMeta"
 
+        -- uMetaVar: unify two type variables
         -- meta variable meets skolem 
         -- => just update
         uMetaVar swapped tv1 (MetaTv _ ref) tv2 (SkolemTv _) cotv
           = do { checkUpdateMeta swapped tv1 ref (mkTyVarTy tv2)
-              ; writeMetaTyVar cotv (mkTyVarTy tv2)
-              ; return ([], True)
-              }
+               ; writeMetaTyVar cotv (mkTyVarTy tv2)
+               ; return Nothing
+               }
 
         -- meta variable meets meta variable 
         -- => be clever about which of the two to update 
@@ -992,8 +1229,8 @@ unifyMetaRule insts
               -- The "nicer to" part only applies if the two kinds are the same,
               -- so we can choose which to do.
 
-              ; writeMetaTyVar cotv (mkTyVarTy tv2)
-              ; return ([], True)
+               ; writeMetaTyVar cotv (mkTyVarTy tv2)
+               ; return Nothing
                }
           where
                 -- Kinds should be guaranteed ok at this point
@@ -1019,130 +1256,6 @@ unifyMetaRule insts
 
 %************************************************************************
 %*                                                                     *
-\section{Normalisation of Insts}
-%*                                                                     *
-%************************************************************************
-
-Normalises a set of dictionaries relative to a set of given equalities (which
-are interpreted as rewrite rules).  We only consider given equalities of the
-form
-
-  F ts ~ t    or    a ~ t
-
-where F is a type family.
-
-\begin{code}
-substEqInDictInsts :: [Inst]    -- given equalities (used as rewrite rules)
-                   -> [Inst]    -- dictinaries to be normalised
-                   -> TcM ([Inst], TcDictBinds)
-substEqInDictInsts eqInsts dictInsts 
-  = do { traceTc (text "substEqInDictInst <-" <+> ppr dictInsts)
-       ; dictInsts' <- 
-           foldlM rewriteWithOneEquality (dictInsts, emptyBag) eqInsts
-       ; traceTc (text "substEqInDictInst ->" <+> ppr dictInsts')
-       ; return dictInsts'
-       }
-  where
-      -- (1) Given equality of form 'F ts ~ t' or 'a ~ t': use for rewriting
-    rewriteWithOneEquality (dictInsts, dictBinds)
-                           eqInst@(EqInst {tci_left  = pattern, 
-                                           tci_right = target})
-      | isOpenSynTyConApp pattern || isTyVarTy pattern
-      = do { (dictInsts', moreDictBinds) <- 
-               genericNormaliseInsts True {- wanted -} applyThisEq dictInsts
-           ; return (dictInsts', dictBinds `unionBags` moreDictBinds)
-           }
-      where
-        applyThisEq = tcGenericNormaliseFamInstPred (return . matchResult)
-
-        -- rewrite in case of an exact match
-        matchResult ty | tcEqType pattern ty = Just (target, eqInstType eqInst)
-                       | otherwise           = Nothing
-
-      -- (2) Given equality has the wrong form: ignore
-    rewriteWithOneEquality (dictInsts, dictBinds) _not_a_rewrite_rule
-      = return (dictInsts, dictBinds)
-\end{code}
-
-
-Take a bunch of Insts (not EqInsts), and normalise them wrt the top-level
-type-function equations, where
-
-       (norm_insts, binds) = normaliseInsts is_wanted insts
-
-If 'is_wanted'
-  = True,  (binds + norm_insts) defines insts       (wanteds)
-  = False, (binds + insts)      defines norm_insts  (givens)
-
-Ie, in the case of normalising wanted dictionaries, we use the normalised
-dictionaries to define the originally wanted ones.  However, in the case of
-given dictionaries, we use the originally given ones to define the normalised
-ones. 
-
-\begin{code}
-normaliseInsts :: Bool                         -- True <=> wanted insts
-              -> [Inst]                        -- wanted or given insts 
-              -> TcM ([Inst], TcDictBinds)     -- normalised insts and bindings
-normaliseInsts isWanted insts 
-  = genericNormaliseInsts isWanted tcNormaliseFamInstPred insts
-
-genericNormaliseInsts  :: Bool                     -- True <=> wanted insts
-                      -> (TcPredType -> TcM (CoercionI, TcPredType))  
-                                                    -- how to normalise
-                      -> [Inst]                    -- wanted or given insts 
-                      -> TcM ([Inst], TcDictBinds) -- normalised insts & binds
-genericNormaliseInsts isWanted fun insts
-  = do { (insts', binds) <- mapAndUnzipM (normaliseOneInst isWanted fun) insts
-       ; return (insts', unionManyBags binds)
-       }
-  where
-    normaliseOneInst isWanted fun
-                    dict@(Dict {tci_pred = pred,
-                                 tci_loc  = loc})
-      = do { traceTc $ text "genericNormaliseInst <-" <+> ppr dict
-          ; (coi, pred') <- fun pred
-
-          ; case coi of
-              IdCo   -> 
-                 do { traceTc $ text "genericNormaliseInst ->" <+> ppr dict
-                    ; return (dict, emptyBag)
-                    }
-                         -- don't use pred' in this case; otherwise, we get
-                         -- more unfolded closed type synonyms in error messages
-              ACo co -> 
-                 do { -- an inst for the new pred
-                   ; dict' <- newDictBndr loc pred'
-                     -- relate the old inst to the new one
-                     -- target_dict = source_dict `cast` st_co
-                   ; let (target_dict, source_dict, st_co) 
-                           | isWanted  = (dict,  dict', mkSymCoercion co)
-                           | otherwise = (dict', dict,  co)
-                             -- we have
-                              --   co :: dict ~ dict'
-                             -- hence, if isWanted
-                             --          dict  = dict' `cast` sym co
-                             --        else
-                             --          dict' = dict  `cast` co
-                         expr      = HsVar $ instToId source_dict
-                         cast_expr = HsWrap (WpCo st_co) expr
-                         rhs       = L (instLocSpan loc) cast_expr
-                         binds     = instToDictBind target_dict rhs
-                     -- return the new inst
-                   ; traceTc $ text "genericNormaliseInst ->" <+> ppr dict'
-                    ; return (dict', binds)
-                   }
-          }
-       
-       -- TOMDO: What do we have to do about ImplicInst, Method, and LitInst??
-    normaliseOneInst _isWanted _fun inst
-      = do { inst' <- zonkInst inst
-          ; return (inst', emptyBag)
-          }
-\end{code}
-
-
-%************************************************************************
-%*                                                                     *
 \section{Errors}
 %*                                                                     *
 %************************************************************************
@@ -1154,45 +1267,64 @@ somethingdifferent message.
 eqInstMisMatch :: Inst -> TcM a
 eqInstMisMatch inst
   = ASSERT( isEqInst inst )
-    do { (env, msg) <- misMatchMsg ty_act ty_exp
-       ; setErrCtxt ctxt $
-           failWithTcM (env, msg)
-       }
+    setErrCtxt ctxt $ failWithMisMatch ty_act ty_exp
   where
-    ty_act           = eqInstLeftTy  inst
-    ty_exp           = eqInstRightTy inst
-    InstLoc _ _ ctxt = instLoc       inst
+    (ty_act, ty_exp) = eqInstTys inst
+    InstLoc _ _ ctxt = instLoc   inst
 
 -----------------------
-misMatchMsg :: TcType -> TcType -> TcM (TidyEnv, SDoc)
+failWithMisMatch :: TcType -> TcType -> TcM a
 -- Generate the message when two types fail to match,
 -- going to some trouble to make it helpful.
 -- The argument order is: actual type, expected type
-misMatchMsg ty_act ty_exp
+failWithMisMatch ty_act ty_exp
   = do { env0 <- tcInitTidyEnv
         ; ty_exp <- zonkTcType ty_exp
         ; ty_act <- zonkTcType ty_act
-       ; (env1, pp_exp, extra_exp) <- ppr_ty env0 ty_exp
-       ; (env2, pp_act, extra_act) <- ppr_ty env1 ty_act
-       ; return (env2, 
-                  sep [sep [ptext SLIT("Couldn't match expected type") <+> pp_exp, 
-                           nest 7 $
-                              ptext SLIT("against inferred type") <+> pp_act],
-                      nest 2 (extra_exp $$ extra_act)]) }
-
-ppr_ty :: TidyEnv -> TcType -> TcM (TidyEnv, SDoc, SDoc)
-ppr_ty env ty
-  = do { let (env1, tidy_ty) = tidyOpenType env ty
-       ; (env2, extra) <- ppr_extra env1 tidy_ty
-       ; return (env2, quotes (ppr tidy_ty), extra) }
-
--- (ppr_extra env ty) shows extra info about 'ty'
-ppr_extra :: TidyEnv -> Type -> TcM (TidyEnv, SDoc)
-ppr_extra env (TyVarTy tv)
-  | isSkolemTyVar tv || isSigTyVar tv
-  = return (env1, pprSkolTvBinding tv1)
+        ; failWithTcM (misMatchMsg env0 (ty_act, ty_exp))
+       }
+
+misMatchMsg :: TidyEnv -> (TcType, TcType) -> (TidyEnv, SDoc)
+misMatchMsg env0 (ty_act, ty_exp)
+  = let (env1, pp_exp, extra_exp) = ppr_ty env0 ty_exp
+       (env2, pp_act, extra_act) = ppr_ty env1 ty_act
+        msg = sep [sep [ptext (sLit "Couldn't match expected type") <+> pp_exp, 
+                       nest 7 $
+                              ptext (sLit "against inferred type") <+> pp_act],
+                  nest 2 (extra_exp $$ extra_act)]
+    in
+    (env2, msg)
+
   where
-    (env1, tv1) = tidySkolemTyVar env tv
+    ppr_ty :: TidyEnv -> TcType -> (TidyEnv, SDoc, SDoc)
+    ppr_ty env ty
+      = let (env1, tidy_ty) = tidyOpenType env ty
+           (env2, extra)  = ppr_extra env1 tidy_ty
+       in
+       (env2, quotes (ppr tidy_ty), extra)
+
+    -- (ppr_extra env ty) shows extra info about 'ty'
+    ppr_extra :: TidyEnv -> Type -> (TidyEnv, SDoc)
+    ppr_extra env (TyVarTy tv)
+      | isTcTyVar tv && (isSkolemTyVar tv || isSigTyVar tv) && not (isUnk tv)
+      = (env1, pprSkolTvBinding tv1)
+      where
+        (env1, tv1) = tidySkolemTyVar env tv
+
+    ppr_extra env _ty = (env, empty)           -- Normal case
+\end{code}
 
-ppr_extra env _ty = return (env, empty)                -- Normal case
+Warn of loopy local equalities that were dropped.
+
+\begin{code}
+warnDroppingLoopyEquality :: TcType -> TcType -> TcM ()
+warnDroppingLoopyEquality ty1 ty2 
+  = do { env0 <- tcInitTidyEnv
+       ; ty1 <- zonkTcType ty1
+       ; ty2 <- zonkTcType ty2
+       ; let (env1 , tidy_ty1) = tidyOpenType env0 ty1
+            (_env2, tidy_ty2) = tidyOpenType env1 ty2
+       ; addWarnTc $ hang (ptext (sLit "Dropping loopy given equality"))
+                      2 (quotes (ppr tidy_ty1 <+> text "~" <+> ppr tidy_ty2))
+       }
 \end{code}