split HaskProofCategory into two files
[coq-hetmet.git] / src / ExtractionMain.v
index 61993a1..4e5a024 100644 (file)
@@ -13,7 +13,6 @@ Require Import Preamble.
 Require Import General.
 
 Require Import NaturalDeduction.
-Require Import NaturalDeductionToLatex.
 
 Require Import HaskKinds.
 Require Import HaskLiteralsAndTyCons.
@@ -36,11 +35,10 @@ Require Import HaskProofToStrong.
 
 Require Import ProgrammingLanguage.
 
-Require Import HaskProofCategory.
-(*
-Require Import HaskStrongCategory.
-*)
-Require Import ReificationsEquivalentToGeneralizedArrows.
+Require Import HaskProofStratified.
+Require Import HaskProofFlattener.
+
+Require Import ReificationsIsomorphicToGeneralizedArrows.
 
 Open Scope string_scope.
 Extraction Language Haskell.
@@ -61,8 +59,7 @@ Extract Inductive unit    => "()" [ "()" ].
 Extract Inlined Constant string_dec => "(==)".
 Extract Inlined Constant ascii_dec => "(==)".
 
-(* adapted from ExtrOcamlString.v *)
-Extract Inductive ascii => "Char" [ "bin2ascii" ] "bin2ascii'".
+Extract Inductive ascii => "Char" [ "you_forgot_to_patch_coq" ] "you_forgot_to_patch_coq".
 Extract Constant zero   => "'\000'".
 Extract Constant one    => "'\001'".
 Extract Constant shift  => "shiftAscii".
@@ -100,7 +97,8 @@ Section core2proof.
   Definition ξ (cv:CoreVar) : LeveledHaskType Γ ★ :=
     match coreVarToWeakVar cv with
       | WExprVar wev => match weakTypeToTypeOfKind φ wev ★ with
-                          | Error s => Prelude_error ("Error in top-level xi: " +++ s)
+                          | Error s => Prelude_error ("Error converting weakType of top-level variable "+++
+                                                         toString cv+++": " +++ s)
                           | OK    t => t @@ nil
                         end
       | WTypeVar _   => Prelude_error "top-level xi got a type variable"
@@ -108,8 +106,6 @@ Section core2proof.
     end.
 
 
-  (* core-to-string (-dcoqpass) *)
-
   Definition header : string :=
     "\documentclass[9pt]{article}"+++eol+++
     "\usepackage{amsmath}"+++eol+++
@@ -128,7 +124,7 @@ Section core2proof.
     eol+++"\end{document}"+++
     eol.
 
-
+  (* core-to-string (-dcoqpass) *)
   Definition coreToStringExpr' (ce:@CoreExpr CoreVar) : ???string :=
     addErrorMessage ("input CoreSyn: " +++ toString ce)
     (addErrorMessage ("input CoreType: " +++ toString (coreTypeOfCoreExpr ce)) (
@@ -227,51 +223,6 @@ Section core2proof.
       apply t.
       Defined.
 
-    Definition env := ★::nil.
-    Definition freshTV : HaskType env ★ := haskTyVarToType (FreshHaskTyVar _).
-    Definition idproof0 : ND Rule [] [env > nil > [] |- [freshTV--->freshTV @@ nil]].
-      eapply nd_comp.
-      eapply nd_comp.
-      eapply nd_rule.
-      apply RVar.
-      eapply nd_rule.
-      eapply (RURule _ _ _ _ (RuCanL _ _)) .
-      eapply nd_rule.
-      eapply RLam.
-      Defined.
-
-(*
-    Definition coreToCoreExpr' (ce:@CoreExpr CoreVar) : ???(@CoreExpr CoreVar) :=
-    addErrorMessage ("input CoreSyn: " +++ toString ce)
-    (addErrorMessage ("input CoreType: " +++ toString (coreTypeOfCoreExpr ce)) (
-      coreExprToWeakExpr ce >>= fun we =>
-        addErrorMessage ("WeakExpr: " +++ toString we)
-          ((addErrorMessage ("CoreType of WeakExpr: " +++ toString (coreTypeOfCoreExpr (weakExprToCoreExpr we)))
-            ((weakTypeOfWeakExpr we) >>= fun t =>
-              (addErrorMessage ("WeakType: " +++ toString t)
-                ((weakTypeToTypeOfKind φ t ★) >>= fun τ =>
-                  addErrorMessage ("HaskType: " +++ toString τ)
-                  ((weakExprToStrongExpr Γ Δ φ ψ ξ (fun _ => true) τ nil we) >>= fun e =>
-                        (let haskProof := @expr2proof _ _ _ _ _ _ e
-                          in (* insert HaskProof-to-HaskProof manipulations here *)
-                   (unFresh (@proof2expr nat _ FreshNat _ _ _ _ (fun _ => Prelude_error "unbound unique") _ haskProof))
-                  >>= fun e' => Error (@toString _ (ExprToString _ _ _ _) (projT2 e'))
-(*
-                  >>= fun e' =>
-                    Prelude_error (@toString _ (@ExprToString nat _ _ _ _ _ _) (projT2 e'))
-  *)                  
-)
-)))))))).
-(*                    Error "X").*)
-(*
-                   strongExprToWeakExpr hetmet_brak hetmet_esc mkWeakTypeVar mkWeakCoerVar mkWeakExprVar uniqueSupply
-                   (projT2 e')
-                         INil
-                         >>= fun q => Error (toString q)
-                  ))))))))).
-*)*)
-
-
     Definition coreToCoreExpr' (ce:@CoreExpr CoreVar) : ???(@CoreExpr CoreVar) :=
       addErrorMessage ("input CoreSyn: " +++ toString ce)
       (addErrorMessage ("input CoreType: " +++ toString (coreTypeOfCoreExpr ce)) (