X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskCoreToWeak.v;h=abcd6b8babfe78affc8d08d5645e53d2532f1765;hp=36399a06135585fdc91a9daef0a886db701f230f;hb=6c953e094065d487e85635df7fd5389271e4a279;hpb=94d7c55025f5df750ce213172c5d2441b5a210e1 diff --git a/src/HaskCoreToWeak.v b/src/HaskCoreToWeak.v index 36399a0..abcd6b8 100644 --- a/src/HaskCoreToWeak.v +++ b/src/HaskCoreToWeak.v @@ -4,8 +4,8 @@ Generalizable All Variables. Require Import Preamble. -Require Import General. Require Import Coq.Lists.List. +Require Import General. Require Import HaskKinds. Require Import HaskLiteralsAndTyCons. Require Import HaskCoreVars. @@ -53,7 +53,7 @@ Fixpoint coreTypeToWeakType' (ct:CoreType) : ???WeakType := | WExprVar _ => Error "encountered expression variable in a modal box type" | WCoerVar _ => Error "encountered coercion variable in a modal box type" end - | _ => Error ("mis-applied modal box tycon: " +++ ct) + | _ => Error ("mis-applied modal box tycon: " +++ toString ct) end else let tc' := if eqd_dec tc ArrowTyCon then WFunTyCon @@ -134,8 +134,8 @@ Fixpoint expectTyConApp (wt:WeakType)(acc:list WeakType) : ???(TyCon * list Weak match wt with | WTyCon tc => OK (tc,acc) | WAppTy t1 t2 => expectTyConApp t1 (t2::acc) - | WTyFunApp tc tys => Error ("expectTyConApp encountered TyFunApp: " +++ wt) - | _ => Error ("expectTyConApp encountered " +++ wt) + | WTyFunApp tc tys => Error ("expectTyConApp encountered TyFunApp: " +++ toString wt) + | _ => Error ("expectTyConApp encountered " +++ toString wt) end. Fixpoint coreExprToWeakExpr (ce:@CoreExpr CoreVar) : ???WeakExpr := @@ -229,9 +229,9 @@ Fixpoint coreExprToWeakExpr (ce:@CoreExpr CoreVar) : ???WeakExpr := | LitAlt lit => OK ((WeakLitAlt lit,nil,nil,nil,e')::rest') | DataAlt dc => let vars := map coreVarToWeakVar vars in OK (((WeakDataAlt dc), - (General.filter (map (fun x => match x with WTypeVar v => Some v | _ => None end) vars)), - (General.filter (map (fun x => match x with WCoerVar v => Some v | _ => None end) vars)), - (General.filter (map (fun x => match x with WExprVar v => Some v | _ => None end) vars)), + (filter (map (fun x => match x with WTypeVar v => Some v | _ => None end) vars)), + (filter (map (fun x => match x with WCoerVar v => Some v | _ => None end) vars)), + (filter (map (fun x => match x with WExprVar v => Some v | _ => None end) vars)), e')::rest') end end) alts)