X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FExtractionMain.v;h=5be52807b489dfd1a977f4c20bc077d361130385;hp=ea89cf42f8eb92e422f424df31fc2c5e10970412;hb=c64ac559ed9448b8aa24abedbc2ad5ca800d1d24;hpb=35d3a59796735e5341389fa6a145f62dcea9c3fc diff --git a/src/ExtractionMain.v b/src/ExtractionMain.v index ea89cf4..5be5280 100644 --- a/src/ExtractionMain.v +++ b/src/ExtractionMain.v @@ -15,7 +15,8 @@ Require Import General. Require Import NaturalDeduction. Require Import HaskKinds. -Require Import HaskLiteralsAndTyCons. +Require Import HaskLiterals. +Require Import HaskTyCons. Require Import HaskCoreVars. Require Import HaskCoreTypes. Require Import HaskCore. @@ -160,7 +161,7 @@ Section core2proof. Definition mkWeakTypeVar (u:Unique)(k:Kind) : WeakTypeVar := weakTypeVar (mkTyVar (mkSystemName u "tv" O) k) k. Definition mkWeakCoerVar (u:Unique)(k:Kind)(t1 t2:WeakType) : WeakCoerVar := - weakCoerVar (mkCoVar (mkSystemName u "cv" O) (weakTypeToCoreType t1) (weakTypeToCoreType t2)) k t1 t2. + weakCoerVar (mkCoVar (mkSystemName u "cv" O) (weakTypeToCoreType t1) (weakTypeToCoreType t2)) t1 t2. Definition mkWeakExprVar (u:Unique)(t:WeakType) : WeakExprVar := weakExprVar (mkExVar (mkSystemName u "ev" O) (weakTypeToCoreType t)) t. @@ -235,12 +236,13 @@ Section core2proof. ND Rule [ Γ > Δ > Σ |- [a ---> s @@ lev ] ] [ Γ > Δ > Σ,,[a @@ lev] |- [ s @@ lev ] ]. - eapply nd_comp; [ idtac | eapply nd_rule; apply (@RApp Γ Δ Σ [a@@lev] a s lev) ]. - eapply nd_comp; [ apply nd_rlecnac | idtac ]. - apply nd_prod. - apply nd_id. + eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply RExch ]. + eapply nd_comp; [ idtac | eapply nd_rule; apply (@RApp Γ Δ [a@@lev] Σ a s lev) ]. + eapply nd_comp; [ apply nd_llecnac | idtac ]. + apply nd_prod. apply nd_rule. apply RVar. + apply nd_id. Defined. Definition fToC1 {Γ}{Δ}{a}{s}{lev} : @@ -279,6 +281,9 @@ Section core2proof. (uniqueSupply : UniqSupply) (lbinds:list (@CoreBind CoreVar)) (hetmet_PGArrowTyCon : TyFun) + (hetmet_PGArrow_unit_TyCon : TyFun) + (hetmet_PGArrow_tensor_TyCon : TyFun) + (hetmet_PGArrow_exponent_TyCon : TyFun) (hetmet_pga_id : CoreVar) (hetmet_pga_comp : CoreVar) (hetmet_pga_first : CoreVar) @@ -298,13 +303,24 @@ Section core2proof. (hetmet_pga_curryr : CoreVar) . - Definition ga_unit TV : RawHaskType TV ★ := @TyFunApp TV UnitTyCon nil ★ TyFunApp_nil. - Definition ga_prod TV (a b:RawHaskType TV ★) : RawHaskType TV ★ := - TApp (TApp (@TyFunApp TV PairTyCon nil _ TyFunApp_nil) a) b. + + Definition ga_unit TV (ec:RawHaskType TV ECKind) : RawHaskType TV ★ := + @TyFunApp TV hetmet_PGArrow_unit_TyCon (ECKind::nil) ★ (TyFunApp_cons _ _ ec TyFunApp_nil). + + Definition ga_prod TV (ec:RawHaskType TV ECKind) (a b:RawHaskType TV ★) : RawHaskType TV ★ := + (@TyFunApp TV + hetmet_PGArrow_tensor_TyCon + (ECKind::★ ::★ ::nil) ★ + (TyFunApp_cons _ _ ec + (TyFunApp_cons _ _ a + (TyFunApp_cons _ _ b + TyFunApp_nil)))). + Definition ga_type {TV}(a:RawHaskType TV ECKind)(b c:RawHaskType TV ★) : RawHaskType TV ★ := TApp (TApp (TApp (@TyFunApp TV hetmet_PGArrowTyCon nil _ TyFunApp_nil) a) b) c. + Definition ga := @ga_mk ga_unit ga_prod (@ga_type). Definition ga_type' {Γ}(a:HaskType Γ ECKind)(b c:HaskType Γ ★) : HaskType Γ ★ := @@ -369,22 +385,22 @@ Section core2proof. |} (ICons _ _ x (ICons _ _ y (ICons _ _ z (ICons _ _ q INil))))). Defined. - Definition gat {Γ}(x:Tree ??(HaskType Γ ★)) := @ga_mk_tree ga_unit ga_prod _ x. + Definition gat {Γ} ec (x:Tree ??(HaskType Γ ★)) := @ga_mk_tree ga_unit ga_prod _ ec x. Instance my_ga : garrow ga_unit ga_prod (@ga_type) := - { ga_id := fun Γ Δ ec l a => mkGlob2 hetmet_pga_id (fun ec a => ga_type' ec a a) ec (gat a) - ; ga_cancelr := fun Γ Δ ec l a => mkGlob2 hetmet_pga_cancelr (fun ec a => ga_type' ec _ a) ec (gat a) - ; ga_cancell := fun Γ Δ ec l a => mkGlob2 hetmet_pga_cancell (fun ec a => ga_type' ec _ a) ec (gat a) - ; ga_uncancelr := fun Γ Δ ec l a => mkGlob2 hetmet_pga_uncancelr (fun ec a => ga_type' ec a _) ec (gat a) - ; ga_uncancell := fun Γ Δ ec l a => mkGlob2 hetmet_pga_uncancell (fun ec a => ga_type' ec a _) ec (gat a) - ; ga_assoc := fun Γ Δ ec l a b c => mkGlob4 hetmet_pga_assoc (fun ec a b c => ga_type' ec _ _) ec (gat a) (gat b) (gat c) - ; ga_unassoc := fun Γ Δ ec l a b c => mkGlob4 hetmet_pga_unassoc (fun ec a b c => ga_type' ec _ _) ec (gat a) (gat b) (gat c) - ; ga_swap := fun Γ Δ ec l a b => mkGlob3 hetmet_pga_swap (fun ec a b => ga_type' ec _ _) ec (gat a) (gat b) - ; ga_drop := fun Γ Δ ec l a => mkGlob2 hetmet_pga_drop (fun ec a => ga_type' ec _ _) ec (gat a) - ; ga_copy := fun Γ Δ ec l a => mkGlob2 hetmet_pga_copy (fun ec a => ga_type' ec _ _) ec (gat a) - ; ga_first := fun Γ Δ ec l a b x => fToC1 (mkGlob4 hetmet_pga_first (fun ec a b c => _) ec (gat a) (gat b) (gat x)) - ; ga_second := fun Γ Δ ec l a b x => fToC1 (mkGlob4 hetmet_pga_second (fun ec a b c => _) ec (gat a) (gat b) (gat x)) - ; ga_comp := fun Γ Δ ec l a b c => fToC2 (mkGlob4 hetmet_pga_comp (fun ec a b c => _) ec (gat a) (gat b) (gat c)) + { ga_id := fun Γ Δ ec l a => mkGlob2 hetmet_pga_id (fun ec a => ga_type' ec a a) ec (gat ec a) + ; ga_cancelr := fun Γ Δ ec l a => mkGlob2 hetmet_pga_cancelr (fun ec a => ga_type' ec _ a) ec (gat ec a) + ; ga_cancell := fun Γ Δ ec l a => mkGlob2 hetmet_pga_cancell (fun ec a => ga_type' ec _ a) ec (gat ec a) + ; ga_uncancelr := fun Γ Δ ec l a => mkGlob2 hetmet_pga_uncancelr (fun ec a => ga_type' ec a _) ec (gat ec a) + ; ga_uncancell := fun Γ Δ ec l a => mkGlob2 hetmet_pga_uncancell (fun ec a => ga_type' ec a _) ec (gat ec a) + ; ga_assoc := fun Γ Δ ec l a b c => mkGlob4 hetmet_pga_assoc (fun ec a b c => ga_type' ec _ _) ec (gat ec a) (gat ec b) (gat ec c) + ; ga_unassoc := fun Γ Δ ec l a b c => mkGlob4 hetmet_pga_unassoc (fun ec a b c => ga_type' ec _ _) ec (gat ec a) (gat ec b) (gat ec c) + ; ga_swap := fun Γ Δ ec l a b => mkGlob3 hetmet_pga_swap (fun ec a b => ga_type' ec _ _) ec (gat ec a) (gat ec b) + ; ga_drop := fun Γ Δ ec l a => mkGlob2 hetmet_pga_drop (fun ec a => ga_type' ec _ _) ec (gat ec a) + ; ga_copy := fun Γ Δ ec l a => mkGlob2 hetmet_pga_copy (fun ec a => ga_type' ec _ _) ec (gat ec a) + ; ga_first := fun Γ Δ ec l a b x => fToC1 (mkGlob4 hetmet_pga_first (fun ec a b c => _) ec (gat ec a) (gat ec b) (gat ec x)) + ; ga_second := fun Γ Δ ec l a b x => fToC1 (mkGlob4 hetmet_pga_second (fun ec a b c => _) ec (gat ec a) (gat ec b) (gat ec x)) + ; ga_comp := fun Γ Δ ec l a b c => fToC2 (mkGlob4 hetmet_pga_comp (fun ec a b c => _) ec (gat ec a) (gat ec b) (gat ec c)) (* ; ga_lit := fun Γ Δ ec l a => nd_rule (RGlobal _ _ _ _ (coreVarToWeakExprVarOrError hetmet_pga_lit))*) (* ; ga_curry := fun Γ Δ ec l a => nd_rule (RGlobal _ _ _ _ (coreVarToWeakExprVarOrError hetmet_pga_curry))*) (* ; ga_apply := fun Γ Δ ec l a => nd_rule (RGlobal _ _ _ _ (coreVarToWeakExprVarOrError hetmet_pga_apply))*) @@ -414,7 +430,7 @@ Section core2proof. ((weakExprToStrongExpr Γ Δ φ ψ ξ (fun _ => true) τ nil we) >>= fun e => (addErrorMessage ("HaskStrong...") - (let haskProof := flatten_proof hetmet_flatten' hetmet_unflatten' + (let haskProof := skolemize_and_flatten_proof hetmet_flatten' hetmet_unflatten' hetmet_flattened_id' my_ga (@expr2proof _ _ _ _ _ _ e) in (* insert HaskProof-to-HaskProof manipulations here *) OK ((@proof2expr nat _ FreshNat _ _ _ _ (fun _ => Prelude_error "unbound unique") _ haskProof) O) @@ -463,6 +479,9 @@ Section core2proof. (uniqueSupply : UniqSupply) (lbinds:list (@CoreBind CoreVar)) (hetmet_PGArrowTyCon : TyFun) + (hetmet_PGArrow_unit_TyCon : TyFun) + (hetmet_PGArrow_tensor_TyCon : TyFun) + (hetmet_PGArrow_exponent_TyCon : TyFun) (hetmet_pga_id : CoreVar) (hetmet_pga_comp : CoreVar) (hetmet_pga_first : CoreVar) @@ -488,6 +507,9 @@ Section core2proof. hetmet_flattened_id uniqueSupply hetmet_PGArrowTyCon + hetmet_PGArrow_unit_TyCon + hetmet_PGArrow_tensor_TyCon +(* hetmet_PGArrow_exponent_TyCon*) hetmet_pga_id hetmet_pga_comp hetmet_pga_first @@ -507,6 +529,7 @@ Section core2proof. hetmet_pga_applyr hetmet_pga_curryl *) + . End core2proof.