better error reporting in coreTypeToWeakType; dont use Prelude_error
[coq-hetmet.git] / src / ExtractionMain.v
index d500e79..2e51d0e 100644 (file)
@@ -92,14 +92,15 @@ Section core2proof.
   (* We need to be able to resolve unbound exprvars, but we can be sure their types will have no
    * free tyvars in them *)
   Definition ξ (cv:CoreVar) : LeveledHaskType Γ ★ :=
-    match coreVarToWeakVar cv with
-      | WExprVar wev => match weakTypeToTypeOfKind φ wev ★ with
+    match coreVarToWeakVar' cv with
+      | OK (WExprVar wev) => match weakTypeToTypeOfKind φ wev ★ with
                           | 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"
-      | WCoerVar _   => Prelude_error "top-level xi got a coercion variable"
+      | OK (WTypeVar _)   => Prelude_error "top-level xi got a type variable"
+      | OK (WCoerVar _)   => Prelude_error "top-level xi got a coercion variable"
+      | Error s           => Prelude_error s
     end.
 
   Definition header : string :=
@@ -228,8 +229,9 @@ Section core2proof.
   End CoreToCore.
 
   Definition coreVarToWeakExprVarOrError cv :=
-    match coreVarToWeakVar cv with
-      | WExprVar wv => wv
+    match addErrorMessage ("in coreVarToWeakExprVarOrError" +++ eol) (coreVarToWeakVar' cv) with
+      | OK (WExprVar wv) => wv
+      | Error s     => Prelude_error s
       | _           => Prelude_error "IMPOSSIBLE"
     end.
 
@@ -237,7 +239,7 @@ Section core2proof.
     ND Rule 
        [ Γ > Δ > Σ             |- [a ---> s  ]@lev ]
        [ Γ > Δ > [a @@ lev],,Σ |-       [ s ]@lev ].
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RExch ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch ].
     eapply nd_comp; [ idtac | eapply nd_rule; eapply RApp ].
     eapply nd_comp; [ apply nd_rlecnac | idtac ].
     apply nd_prod.
@@ -252,7 +254,7 @@ Section core2proof.
     intro pf.
     eapply nd_comp.
     apply pf.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply RCanR ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply ACanR ].
     apply curry.
     Defined.
 
@@ -268,13 +270,15 @@ Section core2proof.
     eapply nd_comp.
     eapply nd_rule.
     eapply RArrange.
-    eapply RCanR.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RExch ].
+    eapply ACanR.
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch ].
     apply curry.
     Defined.
 
   Section coqPassCoreToCore.
     Context
+    (do_flatten : bool)
+    (do_skolemize : bool)
     (hetmet_brak  : CoreVar)
     (hetmet_esc   : CoreVar)
     (hetmet_flatten : CoreVar)
@@ -432,18 +436,41 @@ Section core2proof.
                     ((weakExprToStrongExpr Γ Δ φ ψ ξ (fun _ => true) τ nil we) >>= fun e =>
 
                       (addErrorMessage ("HaskStrong...")
-                        (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 _ _ (flatten_type τ) nil _ (fun _ => Prelude_error "unbound unique") _ haskProof) O)
-                         >>= fun e' =>
-                           (snd e') >>= fun e'' =>
-                         strongExprToWeakExpr hetmet_brak' hetmet_esc'
-                           mkWeakTypeVar mkWeakCoerVar mkWeakExprVar uniqueSupply
-                           (projT2 e'') INil
-                         >>= fun q =>
-                           OK (weakExprToCoreExpr q)
-                    )))))))))).
+                        (if do_skolemize
+                        then
+                             (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 _ _ (flatten_type τ) nil _
+                                (fun _ => Prelude_error "unbound unique") _ haskProof) O)
+                              >>= fun e' => (snd e') >>= fun e'' =>
+                              strongExprToWeakExpr hetmet_brak' hetmet_esc'
+                                mkWeakTypeVar mkWeakCoerVar mkWeakExprVar uniqueSupply
+                                (projT2 e'') INil
+                              >>= fun q => OK (weakExprToCoreExpr q))
+                        else (if do_flatten
+                        then
+                          (let haskProof := flatten_proof (*hetmet_flatten' hetmet_unflatten'
+                                                 hetmet_flattened_id' my_ga*) (@expr2proof _ _ _ _ _ _ _ e)
+                              in (* insert HaskProof-to-HaskProof manipulations here *)
+                              OK ((@proof2expr nat _ FreshNat _ _ τ nil _
+                                (fun _ => Prelude_error "unbound unique") _ haskProof) O)
+                              >>= fun e' => (snd e') >>= fun e'' =>
+                              strongExprToWeakExpr hetmet_brak' hetmet_esc'
+                                mkWeakTypeVar mkWeakCoerVar mkWeakExprVar uniqueSupply
+                                (projT2 e'') INil
+                              >>= fun q => OK (weakExprToCoreExpr q))
+                        else
+                          (let haskProof := @expr2proof _ _ _ _ _ _ _ e
+                              in (* insert HaskProof-to-HaskProof manipulations here *)
+                              OK ((@proof2expr nat _ FreshNat _ _ τ nil _
+                                (fun _ => Prelude_error "unbound unique") _ haskProof) O)
+                              >>= fun e' => (snd e') >>= fun e'' =>
+                              strongExprToWeakExpr hetmet_brak' hetmet_esc'
+                                mkWeakTypeVar mkWeakCoerVar mkWeakExprVar uniqueSupply
+                                (projT2 e'') INil
+                              >>= fun q => OK (weakExprToCoreExpr q))))
+                  ))))))))).
 
     Definition coreToCoreExpr (ce:@CoreExpr CoreVar) : (@CoreExpr CoreVar) :=
       match coreToCoreExpr' ce with
@@ -473,6 +500,8 @@ Section core2proof.
   End coqPassCoreToCore.
 
     Definition coqPassCoreToCore 
+    (do_flatten  : bool)
+    (do_skolemize  : bool)
     (hetmet_brak  : CoreVar)
     (hetmet_esc   : CoreVar)
     (hetmet_flatten   : CoreVar)
@@ -502,6 +531,8 @@ Section core2proof.
     (hetmet_pga_curryl : CoreVar)
     (hetmet_pga_curryr : CoreVar) : list (@CoreBind CoreVar) :=
     coqPassCoreToCore'
+       do_flatten
+       do_skolemize
        hetmet_brak  
        hetmet_esc   
        hetmet_flatten