X-Git-Url: http://git.megacz.com/?a=blobdiff_plain;f=src%2FHaskStrongToWeak.v;h=191162b3824c1e1e0eb127e66400d329684a89f3;hb=8efffc7368b5e54c42461f45a9708ff2828409a4;hp=a7f6b777921e9aa9efafe3e0f3265fc012939cb3;hpb=1f411b48dd607e76a65903e8506d0ae5e7470321;p=coq-hetmet.git diff --git a/src/HaskStrongToWeak.v b/src/HaskStrongToWeak.v index a7f6b77..191162b 100644 --- a/src/HaskStrongToWeak.v +++ b/src/HaskStrongToWeak.v @@ -12,6 +12,8 @@ Require Import Coq.Init.Specif. Require Import HaskKinds. Require Import HaskCoreLiterals. Require Import HaskCoreVars. +Require Import HaskCoreTypes. +Require Import HaskCore. Require Import HaskWeakTypes. Require Import HaskWeakVars. Require Import HaskWeak. @@ -56,11 +58,6 @@ Definition typeToWeakType (f:Fresh Kind (fun _ => WeakTypeVar)) {Γ}{κ}(τ:HaskType Γ κ)(φ:InstantiatedTypeEnv (fun _ => WeakTypeVar) Γ) : WeakType := rawHaskTypeToWeakType f (τ _ φ). -Variable unsafeCoerce : WeakCoercion. Extract Inlined Constant unsafeCoerce => "Coercion.unsafeCoerce". - -Definition strongCoercionToWeakCoercion {Γ}{Δ}{ck}(τ:HaskCoercion Γ Δ ck)(φ:InstantiatedTypeEnv (fun _ => WeakTypeVar) Γ) - : WeakCoercion := unsafeCoerce. - (* This can be used to turn HaskStrong's with arbitrary expression variables into HaskStrong's which use WeakExprVar's * for their variables; those can subsequently be passed to the following function (strongExprToWeakExpr) *) (* @@ -72,6 +69,12 @@ Definition reindexStrongExpr {VV}{HH}{eqVV:EqDecidable VV}{eqHH:EqDecidable HH} Definition updateITE {Γ:TypeEnv}{TV:Kind->Type}{κ}(tv:TV κ)(ite:InstantiatedTypeEnv TV Γ) : InstantiatedTypeEnv TV (κ::Γ) := tv::::ite. +Section strongExprToWeakExpr. + + Context (hetmet_brak : CoreVar). + Context (hetmet_esc : CoreVar). +Axiom FIXME : forall {t}, t. + Fixpoint strongExprToWeakExpr {Γ}{Δ}{ξ}{τ} (ftv:Fresh Kind (fun _ => WeakTypeVar)) (fcv:Fresh (WeakType*WeakType) (fun _ => WeakCoerVar)) @@ -83,15 +86,19 @@ match exp as E in @Expr _ _ G D X L return InstantiatedTypeEnv (fun _ => WeakTyp | EApp Γ' _ _ _ _ _ e1 e2 => fun ite => WEApp (strongExprToWeakExpr ftv fcv e1 ite) (strongExprToWeakExpr ftv fcv e2 ite) | ELam Γ' _ _ _ t _ cv e => fun ite => WELam (weakExprVar cv (typeToWeakType ftv t ite)) (strongExprToWeakExpr ftv fcv e ite) | ELet Γ' _ _ t _ _ ev e1 e2 => fun ite => WELet (weakExprVar ev (typeToWeakType ftv t ite)) (strongExprToWeakExpr ftv fcv e1 ite) (strongExprToWeakExpr ftv fcv e2 ite) -| EEsc Γ' _ _ ec t _ e => fun ite => WEEsc (ec _ ite) (strongExprToWeakExpr ftv fcv e ite) (typeToWeakType ftv t ite) -| EBrak Γ' _ _ ec t _ e => fun ite => WEBrak (ec _ ite) (strongExprToWeakExpr ftv fcv e ite) (typeToWeakType ftv t ite) -| ECast Γ Δ ξ t1 t2 γ l e => fun ite => WECast (strongExprToWeakExpr ftv fcv e ite) (strongCoercionToWeakCoercion γ ite) +| EEsc Γ' _ _ ec t _ e => fun ite => WEEsc hetmet_esc (ec _ ite) (strongExprToWeakExpr ftv fcv e ite) (typeToWeakType ftv t ite) +| EBrak Γ' _ _ ec t _ e => fun ite => WEBrak hetmet_brak (ec _ ite) (strongExprToWeakExpr ftv fcv e ite) (typeToWeakType ftv t ite) | ENote _ _ _ _ n e => fun ite => WENote n (strongExprToWeakExpr ftv fcv e ite) | ETyApp Γ Δ κ σ τ ξ l e => fun ite => WETyApp (strongExprToWeakExpr ftv fcv e ite) (typeToWeakType ftv τ ite) | ELetRec _ _ _ _ _ vars elrb e =>fun ite=>WELetRec (exprLetRec2WeakExprLetRec ftv fcv elrb ite)(strongExprToWeakExpr ftv fcv e ite) -| ECoApp Γ Δ κ σ₁ σ₂ γ σ ξ l e => fun ite => WECoApp (strongExprToWeakExpr ftv fcv e ite) (strongCoercionToWeakCoercion γ ite) +| ECast Γ Δ ξ t1 t2 γ l e => fun ite => FIXME +(* WECast (strongExprToWeakExpr ftv fcv e ite) (strongCoercionToWeakCoercion γ ite)*) + +| ECoApp Γ Δ κ σ₁ σ₂ γ σ ξ l e => fun ite => FIXME +(* WECoApp (strongExprToWeakExpr ftv fcv e ite) (strongCoercionToWeakCoercion γ ite)*) | ECase Γ Δ ξ l tc tbranches atypes e alts => fun ite => WECase + FIXME (strongExprToWeakExpr ftv fcv e ite) (@typeToWeakType ftv Γ _ tbranches ite) tc @@ -137,3 +144,4 @@ match elrb as E in ELetRecBindings G D X L V | ELR_leaf _ _ ξ' cv v e => fun ite => [((weakExprVar v (typeToWeakType ftv (unlev (ξ' v)) ite)),(strongExprToWeakExpr ftv fcv e ite))] | ELR_branch _ _ _ _ _ _ b1 b2 => fun ite => (exprLetRec2WeakExprLetRec ftv fcv b1 ite),, (exprLetRec2WeakExprLetRec ftv fcv b2 ite) end. +End strongExprToWeakExpr. \ No newline at end of file