X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskCoreToWeak.v;h=ccd497355f0be02006b2dc272f47ed958c6acd8f;hp=78223cf4d10c87d14e22ea6f1c9b161784220466;hb=38e0c88fa03d930293f980681fa34a667402a20d;hpb=e1c9b2e7a28fabb4c0a9ce3cbd4d2ae099d60b5a diff --git a/src/HaskCoreToWeak.v b/src/HaskCoreToWeak.v index 78223cf..ccd4973 100644 --- a/src/HaskCoreToWeak.v +++ b/src/HaskCoreToWeak.v @@ -29,6 +29,12 @@ Variable hetmet_brak_name : CoreName. Extract Inlined Constant he Variable hetmet_esc_name : CoreName. Extract Inlined Constant hetmet_esc_name => "PrelNames.hetmet_esc_name". Variable hetmet_csp_name : CoreName. Extract Inlined Constant hetmet_csp_name => "PrelNames.hetmet_csp_name". +Definition mkTyFunApplication (tf:TyFun)(lwt:list WeakType) : ???WeakType := + split_list lwt (length (fst (tyFunKind tf))) >>= + fun argsrest => + let (args,rest) := argsrest in + OK (fold_left (fun x y => WAppTy x y) rest (WTyFunApp tf args)). + Fixpoint coreTypeToWeakType' (ct:CoreType) : ???WeakType := match ct with | TyVarTy cv => match coreVarToWeakVar cv with @@ -45,7 +51,7 @@ Fixpoint coreTypeToWeakType' (ct:CoreType) : ???WeakType := | a::b => coreTypeToWeakType' a >>= fun a' => rec b >>= fun b' => OK (a'::b') end) lct) in match tyConOrTyFun tc_ with - | inr tf => recurse >>= fun recurse' => OK (WTyFunApp tf recurse') + | inr tf => recurse >>= fun recurse' => mkTyFunApplication tf recurse' | inl tc => if eqd_dec tc ModalBoxTyCon then match lct with | ((TyVarTy ec)::tbody::nil) =>