]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/components/ng_kernel/nCicExtraction.ml
Capitalization of variables bound in patterns fixed.
[helm.git] / matita / components / ng_kernel / nCicExtraction.ml
index fb98942b6d4a6241c9ffd9732cedf9fb08a1e843..5690d1c9f7f58da6fc55fa69c65dfe73a4ceeb77 100644 (file)
@@ -69,14 +69,14 @@ type typ =
 
 let rec size_of_type =
   function
-    | Var v -> 1
+    | Var _ -> 1
     | Unit -> 1
     | Top -> 1
-    | TConst c -> 1
+    | TConst _ -> 1
     | Arrow _ -> 2
     | TSkip t -> size_of_type t
     | Forall _ -> 2
-    | TAppl l -> 1
+    | TAppl _ -> 1
 ;;
 
 (* None = dropped abstraction *)
@@ -95,7 +95,7 @@ type term =
   | LetIn of string * (* typ **) term * term
   | Match of reference * term * (term_context * term) list
   | BottomElim
-  | TLambda of (* string **) term
+  | TLambda of string * term
   | Inst of (*typ_former **) term
   | Skip of term
   | UnsafeCoerce of term
@@ -104,39 +104,70 @@ type term =
 type term_former_decl = term_context * typ
 type term_former_def = term_former_decl * term
 
+let mk_appl f x =
+ match f with
+    Appl args -> Appl (args@[x])
+  | _ -> Appl [f;x]
+
 let rec size_of_term =
   function
-    | Rel r -> 1
+    | Rel _ -> 1
     | UnitTerm -> 1
-    | Const c -> 1
-    | Lambda (name, body) -> 1 + size_of_term body
+    | Const _ -> 1
+    | Lambda (_, body) -> 1 + size_of_term body
     | Appl l -> List.length l
-    | LetIn (name, def, body) -> 1 + size_of_term def + size_of_term body
-    | Match (name, case, pats) -> 1 + size_of_term case + List.length pats
+    | LetIn (_, def, body) -> 1 + size_of_term def + size_of_term body
+    | Match (_, case, pats) -> 1 + size_of_term case + List.length pats
     | BottomElim -> 1
-    | TLambda t -> size_of_term t
+    | TLambda (_,t) -> size_of_term t
     | Inst t -> size_of_term t
     | Skip t -> size_of_term t
     | UnsafeCoerce t -> 1 + size_of_term t
 ;;
-let unitty =
- NCic.Const (NReference.reference_of_spec (NUri.uri_of_string "cic:/matita/basics/types/unit.ind") (NReference.Ind (true,0,0)));;
+
+module ReferenceMap = Map.Make(struct type t = NReference.reference let compare = NReference.compare end)
+
+type info_el = typ_context * typ option
+type info = (NReference.reference * info_el) list
+type db = info_el ReferenceMap.t
+
+let empty_info = []
+
+let register_info db (ref,info) = ReferenceMap.add ref info db
+
+let register_infos = List.fold_left register_info
 
 type obj_kind =
    TypeDeclaration of typ_former_decl
  | TypeDefinition of typ_former_def
  | TermDeclaration of term_former_decl
  | TermDefinition of term_former_def
- | LetRec of (NReference.reference * obj_kind) list
+ | LetRec of (info * NReference.reference * obj_kind) list
    (* reference, left, right, constructors *)
- | Algebraic of (NReference.reference * typ_context * typ_context * (NReference.reference * typ) list) list
+ | Algebraic of (info * NReference.reference * typ_context * typ_context * (NReference.reference * typ) list) list
 
-type obj = NReference.reference * obj_kind
+type obj = info * NReference.reference * obj_kind
 
 (* For LetRec and Algebraic blocks *)
 let dummyref =
  NReference.reference_of_string "cic:/matita/dummy/dummy.ind(1,1,1)"
 
+type not_term = [`Kind | `KindOrType | `PropKind | `Proposition | `Type];;
+
+let max_not_term t1 t2 =
+ match t1,t2 with
+    `KindOrType,_
+  | _,`KindOrType -> `KindOrType
+  | `Kind,_
+  | _,`Kind -> `Kind
+  | `Type,_
+  | _,`Type -> `Type
+  | `PropKind,_
+  | _,`PropKind -> `PropKind
+  | `Proposition,`Proposition -> `Proposition
+
+let sup = List.fold_left max_not_term `Proposition
+
 let rec classify_not_term status context t =
  match NCicReduction.whd status ~subst:[] context t with
   | NCic.Sort s ->
@@ -150,14 +181,23 @@ let rec classify_not_term status context t =
      classify_not_term status ((b,NCic.Decl s)::context) t
   | NCic.Implicit _
   | NCic.LetIn _
-  | NCic.Lambda _
   | NCic.Const (NReference.Ref (_,NReference.CoFix _))
-  | NCic.Appl [] -> assert false (* NOT POSSIBLE *)
-  | NCic.Match _
+  | NCic.Appl [] ->
+     assert false (* NOT POSSIBLE *)
+  | NCic.Lambda (n,s,t) ->
+     (* Lambdas can me met in branches of matches, expecially when the
+        output type is a product *)
+     classify_not_term status ((n,NCic.Decl s)::context) t
+  | NCic.Match (_,_,_,pl) ->
+     let classified_pl = List.map (classify_not_term status context) pl in
+      sup classified_pl
   | NCic.Const (NReference.Ref (_,NReference.Fix _)) ->
-     (* be aware: we can be the head of an application *)
-     assert false (* TODO *)
+     assert false (* IMPOSSIBLE *)
   | NCic.Meta _ -> assert false (* TODO *)
+  | NCic.Appl (NCic.Const (NReference.Ref (_,NReference.Fix (i,_,_)) as r)::_)->
+     let l,_,_ = NCicEnvironment.get_checked_fixes_or_cofixes status r in
+     let _,_,_,_,bo = List.nth l i in
+      classify_not_term status [] bo
   | NCic.Appl (he::_) -> classify_not_term status context he
   | NCic.Rel n ->
      let rec find_sort typ =
@@ -201,8 +241,6 @@ let rec classify_not_term status context t =
      assert false (* IMPOSSIBLE *)
 ;;
 
-type not_term = [`Kind | `KindOrType | `PropKind | `Proposition | `Type];;
-
 let classify status ~metasenv context t =
  match NCicTypeChecker.typeof status ~metasenv ~subst:[] context t with
   | NCic.Sort _ ->
@@ -262,10 +300,6 @@ let rec skip_args status ~metasenv context =
       | `Term _ -> assert false (* IMPOSSIBLE *)
 ;;
 
-module ReferenceMap = Map.Make(struct type t = NReference.reference let compare = NReference.compare end)
-
-type db = (typ_context * typ option) ReferenceMap.t
-
 class type g_status =
  object
   method extraction_db: db
@@ -323,15 +357,21 @@ let rec split_typ_lambdas status n ~metasenv context typ =
 ;;
 
 
-let context_of_typformer status ~metasenv context =
+let rev_context_of_typformer status ~metasenv context =
  function
     NCic.Const (NReference.Ref (_,NReference.Ind _) as ref)
   | NCic.Const (NReference.Ref (_,NReference.Def _) as ref)
   | NCic.Const (NReference.Ref (_,NReference.Decl) as ref)
   | NCic.Const (NReference.Ref (_,NReference.Fix _) as ref) ->
-     (try fst (ReferenceMap.find ref status#extraction_db)
+     (try List.rev (fst (ReferenceMap.find ref status#extraction_db))
       with
-       Not_found -> assert false (* IMPOSSIBLE *))
+       Not_found ->
+        (* This can happen only when we are processing the type of
+           the constructor of a small singleton type. In this case
+           we are not interested in all the type, but only in its
+           backbone. That is why we can safely return the dummy context here *)
+        let rec dummy = None::dummy in
+        dummy)
   | NCic.Match _ -> assert false (* TODO ???? *)
   | NCic.Rel n ->
      let typ =
@@ -340,7 +380,7 @@ let context_of_typformer status ~metasenv context =
        | _,NCic.Def _ -> assert false (* IMPOSSIBLE *) in
      let typ_ctx = snd (HExtlib.split_nth n context) in
      let typ = kind_of status ~metasenv typ_ctx typ in
-      fst (split_kind_prods [] typ)
+      List.rev (fst (split_kind_prods [] typ))
   | NCic.Meta _ -> assert false (* TODO *)
   | NCic.Const (NReference.Ref (_,NReference.Con _))
   | NCic.Const (NReference.Ref (_,NReference.CoFix _))
@@ -370,22 +410,74 @@ let rec typ_of status ~metasenv context k =
   | NCic.Lambda _ -> assert false (* LAMBDA-LIFT INNER DECLARATION *)
   | NCic.Rel n -> Var n
   | NCic.Const ref -> TConst ref
+  | NCic.Match (_,_,_,_)
+  | NCic.Appl (NCic.Const (NReference.Ref (_,NReference.Fix _))::_) ->
+     Top
   | NCic.Appl (he::args) ->
-     let he_context = context_of_typformer status ~metasenv context he in
+     let rev_he_context= rev_context_of_typformer status ~metasenv context he in
      TAppl (typ_of status ~metasenv context he ::
       List.map
        (function None -> Unit | Some ty -> typ_of status ~metasenv context ty)
-       (skip_args status ~metasenv context (List.rev he_context,args)))
+       (skip_args status ~metasenv context (rev_he_context,args)))
   | NCic.Appl _ -> assert false (* TODO: when head is a match/let rec;
                                    otherwise NOT A TYPE *)
-  | NCic.Meta _
-  | NCic.Match (_,_,_,_) -> assert false (* TODO *)
+  | NCic.Meta _ -> assert false (*TODO*)
+;;
+
+let rec fomega_lift_type_from n k =
+ function
+  | Var m as t -> if m < k then t else Var (m + n)
+  | Top -> Top
+  | TConst _ as t -> t
+  | Unit -> Unit
+  | Arrow (ty1,ty2) -> Arrow (fomega_lift_type_from n k ty1,fomega_lift_type_from n (k+1) ty2)
+  | TSkip t -> TSkip (fomega_lift_type_from n (k+1) t)
+  | Forall (name,kind,t) -> Forall (name,kind,fomega_lift_type_from n (k+1) t)
+  | TAppl args -> TAppl (List.map (fomega_lift_type_from n k) args)
+
+let fomega_lift_type n t =
+ if n = 0 then t else fomega_lift_type_from n 0 t
+
+let fomega_lift_term n t =
+ let rec fomega_lift_term n k =
+  function
+   | Rel m as t -> if m < k then t else Rel (m + n)
+   | BottomElim
+   | UnitTerm
+   | Const _ as t -> t
+   | Lambda (name,t) -> Lambda (name,fomega_lift_term n (k+1) t)
+   | TLambda (name,t) -> TLambda (name,fomega_lift_term n (k+1) t)
+   | Appl args -> Appl (List.map (fomega_lift_term n k) args)
+   | LetIn (name,m,bo) ->
+      LetIn (name, fomega_lift_term n k m, fomega_lift_term n (k+1) bo)
+   | Match (ref,t,pl) ->
+      let lift_p (ctx,t) =
+       let lift_context ctx =
+        let len = List.length ctx in
+         HExtlib.list_mapi
+          (fun el i ->
+            let j = len - i - 1 in
+            match el with
+               None
+             | Some (_,`OfKind  _) as el -> el
+             | Some (name,`OfType t) ->
+                Some (name,`OfType (fomega_lift_type_from n (k+j) t))
+          ) ctx
+       in
+        lift_context ctx, fomega_lift_term n (k + List.length ctx) t
+      in
+      Match (ref,fomega_lift_term n k t,List.map lift_p pl)
+   | Inst t -> Inst (fomega_lift_term n k t)
+   | Skip t -> Skip (fomega_lift_term n (k+1) t)
+   | UnsafeCoerce t -> UnsafeCoerce (fomega_lift_term n k t)
+ in
+  if n = 0 then t else fomega_lift_term n 0 t
 ;;
 
 let rec fomega_subst k t1 =
  function
   | Var n ->
-     if k=n then t1
+     if k=n then fomega_lift_type k t1
      else if n < k then Var n
      else Var (n-1)
   | Top -> Top
@@ -419,7 +511,7 @@ let rec term_of status ~metasenv context =
      (* CSC: non-invariant assumed here about "_" *)
      (match classify status ~metasenv context ty with
        | `Kind ->
-           TLambda (term_of status ~metasenv ((b,NCic.Decl ty)::context) bo)
+           TLambda (b,term_of status ~metasenv ((b,NCic.Decl ty)::context) bo)
        | `KindOrType (* ??? *)
        | `Type ->
            Lambda (b, term_of status ~metasenv ((b,NCic.Decl ty)::context) bo)
@@ -450,21 +542,10 @@ let rec term_of status ~metasenv context =
   | NCic.Appl (he::args) ->
      eat_args status metasenv
       (term_of status ~metasenv context he) context
+      (*BUG: recomputing every time the type of the head*)
       (typ_of status ~metasenv context
         (NCicTypeChecker.typeof status ~metasenv ~subst:[] context he))
       args
-(*
-     let he_context = context_of_typformer status ~metasenv context he in
-     let process_args he =
-      function
-         `Empty -> he
-       | `Inst tl -> Inst (process_args he tl)
-       | `Appl (arg,tl) -> process_args (Appl (he,... arg)) tl
-     in
-     Appl (typ_of status ~metasenv context he ::
-      process_args (typ_of status ~metasenv context he)
-       (skip_term_args status ~metasenv context (List.rev he_context,args))
-*)
   | NCic.Appl _ -> assert false (* TODO: when head is a match/let rec;
                                    otherwise NOT A TYPE *)
   | NCic.Meta _ -> assert false (* TODO *)
@@ -477,20 +558,37 @@ let rec term_of status ~metasenv context =
       in
        let rec eat_branch n ty context ctx pat =
          match (ty, pat) with
-         | NCic.Prod (_, _, t), _ when n > 0 ->
+         | TSkip t,_
+         | Forall (_,_,t),_
+         | Arrow (_, t), _ when n > 0 ->
             eat_branch (pred n) t context ctx pat 
-         | NCic.Prod (_, _, t), NCic.Lambda (name, ty, t') ->
+         | _, _ when n > 0 -> assert false (*BUG: is Top possible here?*)
+         (*CSC: unify the three cases below? *)
+         | Arrow (_, t), NCic.Lambda (name, ty, t') ->
            let ctx =
-            (*BUG: we should classify according to the constructor type*)
-            (Some (name,`OfType (*(typ_of status ~metasenv context ty)*)Unit))::ctx in
+            (Some (name,`OfType (typ_of status ~metasenv context ty)))::ctx in
            let context = (name,NCic.Decl ty)::context in
             eat_branch 0 t context ctx t'
-         (*BUG here, eta-expand!*)
+         | Forall (_,_,t),NCic.Lambda (name, ty, t') ->
+           let ctx =
+            (Some (name,`OfKind (kind_of status ~metasenv context ty)))::ctx in
+           let context = (name,NCic.Decl ty)::context in
+            eat_branch 0 t context ctx t'
+         | TSkip t,NCic.Lambda (name, ty, t') ->
+            let ctx = None::ctx in
+            let context = (name,NCic.Decl ty)::context in
+             eat_branch 0 t context ctx t'
+         | Top,_ -> assert false (* IMPOSSIBLE *)
+         | TSkip _, _
+         | Forall _,_
+         | Arrow _,_ -> assert false (*BUG here, eta-expand!*)
          | _, _ -> context,ctx, pat
        in
         try
          List.map2
-          (fun (_, name, ty) pat ->
+          (fun (_, _name, ty) pat ->
+            (*BUG: recomputing every time the type of the constructor*)
+            let ty = typ_of status ~metasenv context ty in
             let context,lhs,rhs = eat_branch leftno ty context [] pat in
             let rhs =
              (* UnsafeCoerce not always required *)
@@ -507,7 +605,7 @@ let rec term_of status ~metasenv context =
       | `Proposition ->
           (match patterns_of pl with
               [] -> BottomElim
-            | [_lhs,rhs] -> rhs (*BUG HERE: Rels are not ok, bound in the _lhs*)
+            | [lhs,rhs] -> fomega_lift_term (- List.length lhs) rhs
             | _ -> assert false)
       | `Type ->
           Match (ref,term_of status ~metasenv context t, patterns_of pl))
@@ -515,11 +613,6 @@ and eat_args status metasenv acc context tyhe =
  function
     [] -> acc
   | arg::tl ->
-     let mk_appl f x =
-      match f with
-         Appl args -> Appl (args@[x])
-       | _ -> Appl [f;x]
-     in
      match fomega_whd status tyhe with
         Arrow (s,t) ->
          let arg =
@@ -527,12 +620,39 @@ and eat_args status metasenv acc context tyhe =
              Unit -> UnitTerm
            | _ -> term_of status ~metasenv context arg in
          eat_args status metasenv (mk_appl acc arg) context t tl
+      | Top ->
+         let arg =
+          match classify status ~metasenv context arg with
+           | `PropKind
+           | `Proposition
+           | `Term `TypeFormer
+           | `Term `PropFormer
+           | `Term `Proof
+           | `Type
+           | `KindOrType
+           | `Kind -> UnitTerm
+           | `Term `TypeFormerOrTerm
+           | `Term `Term -> term_of status ~metasenv context arg
+         in
+          eat_args status metasenv
+           (UnsafeCoerce (mk_appl acc (UnsafeCoerce arg)))
+           context Top tl
       | Forall (_,_,t) ->
-         eat_args status metasenv (Inst acc)
-          context (fomega_subst 1 (typ_of status ~metasenv context arg) t) tl
+         (match classify status ~metasenv context arg with
+           | `PropKind -> assert false (*TODO: same as below, coercion needed???*)
+           | `Proposition
+           | `Term `TypeFormer
+           | `Kind ->
+               eat_args status metasenv (UnsafeCoerce (Inst acc))
+                context (fomega_subst 1 Unit t) tl
+           | `Term _ -> assert false (*TODO: ????*)
+           | `KindOrType
+           | `Type ->
+               eat_args status metasenv (Inst acc)
+                context (fomega_subst 1 (typ_of status ~metasenv context arg) t)
+                 tl)
       | TSkip t ->
          eat_args status metasenv acc context t tl
-      | Top -> assert false (*TODO: HOW??*)
       | Unit | Var _ | TConst _ | TAppl _ -> assert false (* NOT A PRODUCT *)
 ;;
 
@@ -562,10 +682,10 @@ let object_of_constant status ~metasenv ref bo ty =
                 ctx0 ctx
               in
               let bo = typ_of status ~metasenv ctx bo in
+              let info = ref,(nicectx,Some bo) in
                status#set_extraction_db
-                (ReferenceMap.add ref (nicectx,Some bo)
-                  status#extraction_db),
-               Success (ref,TypeDefinition((nicectx,res),bo))
+                (register_info status#extraction_db info),
+               Success ([info],ref,TypeDefinition((nicectx,res),bo))
           | `Kind -> status, Erased (* DPM: but not really, more a failure! *)
           | `PropKind
           | `Proposition -> status, Erased
@@ -577,7 +697,7 @@ let object_of_constant status ~metasenv ref bo ty =
        (* CSC: TO BE FINISHED, REF NON REGISTERED *)
        let ty = typ_of status ~metasenv [] ty in
         status,
-        Success (ref, TermDefinition (split_typ_prods [] ty, term_of status ~metasenv [] bo))
+        Success ([],ref, TermDefinition (split_typ_prods [] ty, term_of status ~metasenv [] bo))
     | `Term _ -> status, Failure "Non-term classified as a term.  This is a bug."
 ;;
 
@@ -596,9 +716,10 @@ let object_of_inductive status ~metasenv uri ind leftno il =
           let right,left = HExtlib.split_nth (List.length ctx - leftno) ctx in
           let ref =
            NReference.reference_of_spec uri (NReference.Ind (ind,i,leftno)) in
+          let info = ref,(ctx,None) in
           let status =
            status#set_extraction_db
-            (ReferenceMap.add ref (ctx,None) status#extraction_db) in
+            (register_info status#extraction_db info) in
           let cl =
            HExtlib.list_mapi
             (fun (_,_,ty) j ->
@@ -608,12 +729,12 @@ let object_of_inductive status ~metasenv uri ind leftno il =
                NReference.mk_constructor (j+1) ref,ty
             ) cl
           in
-           status,i+1,(ref,left,right,cl)::res
+           status,i+1,([info],ref,left,right,cl)::res
    ) (status,0,[]) il
  in
   match rev_tyl with
      [] -> status,Erased
-   | _ -> status, Success (dummyref, Algebraic (List.rev rev_tyl))
+   | _ -> status, Success ([], dummyref, Algebraic (List.rev rev_tyl))
 ;;
 
 let object_of status (uri,height,metasenv,subst,obj_kind) =
@@ -625,15 +746,16 @@ let object_of status (uri,height,metasenv,subst,obj_kind) =
             | `Kind ->
               let ty = kind_of status ~metasenv [] ty in
               let ctx,_ as res = split_kind_prods [] ty in
+              let info = ref,(ctx,None) in
                 status#set_extraction_db
-                  (ReferenceMap.add ref (ctx,None) status#extraction_db),
-                    Success (ref, TypeDeclaration res)
+                  (register_info status#extraction_db info),
+                    Success ([info],ref, TypeDeclaration res)
             | `PropKind
             | `Proposition -> status, Erased
             | `Type
             | `KindOrType (*???*) ->
               let ty = typ_of status ~metasenv [] ty in
-                status, Success (ref, TermDeclaration (split_typ_prods [] ty))
+                status, Success ([],ref, TermDeclaration (split_typ_prods [] ty))
             | `Term _ -> status, Failure "Type classified as a term.  This is a bug.")
         | NCic.Constant (_,_,Some bo,ty,_) ->
            let ref = NReference.reference_of_spec uri (NReference.Def height) in
@@ -651,10 +773,10 @@ let object_of status (uri,height,metasenv,subst,obj_kind) =
               in
               let status,obj = object_of_constant ~metasenv status ref bo ty in
                 match obj with
-                  | Success (ref,obj) -> i+1,status, (ref,obj)::res
+                  | Success (info,ref,obj) -> i+1,status, (info,ref,obj)::res
                   | _ -> i+1,status, res) fs (0,status,[])
           in
-            status, Success (dummyref,LetRec objs)
+            status, Success ([],dummyref,LetRec objs)
         | NCic.Inductive (ind,leftno,il,_) ->
            object_of_inductive status ~metasenv uri ind leftno il
 
@@ -748,8 +870,8 @@ let classify_reference status ref =
   if ReferenceMap.mem ref status#extraction_db then
     `TypeName
   else
-   let NReference.Ref (_,ref) = ref in
-    match ref with
+   let NReference.Ref (_,r) = ref in
+    match r with
        NReference.Con _ -> `Constructor
      | NReference.Ind _ -> assert false
      | _ -> `FunctionName
@@ -784,7 +906,7 @@ let rec pretty_print_type status ctxt =
   function
     | Var n -> List.nth ctxt (n-1)
     | Unit -> "()"
-    | Top -> assert false (* ??? *)
+    | Top -> "Top"
     | TConst ref -> pp_ref status ref
     | Arrow (t1,t2) ->
         bracket size_of_type (pretty_print_type status ctxt) t1 ^ " -> " ^
@@ -801,18 +923,21 @@ let rec pretty_print_type status ctxt =
    | TAppl tl -> String.concat " " (List.map (pretty_print_type status ctxt) tl)
 
 let pretty_print_term_context status ctx1 ctx2 =
- let name_context, res =
+ let name_context, rev_res =
   List.fold_right
-    (fun el (ctx1,res) ->
+    (fun el (ctx1,rev_res) ->
       match el with
-         None -> ""@::ctx1,res
-       | Some (name,`OfKind _) -> name@::ctx1,res
+         None -> ""@::ctx1,rev_res
+       | Some (name,`OfKind _) ->
+          let name = capitalize `TypeVariable name in
+           name@::ctx1,rev_res
        | Some (name,`OfType typ) ->
+          let name = capitalize `TypeVariable name in
           let name,ctx1 = name@:::ctx1 in
            name::ctx1,
-            ("(" ^ name ^ " :: " ^ pretty_print_type status ctx1 typ ^ ")")::res
+            ("(" ^ name ^ " :: " ^ pretty_print_type status ctx1 typ ^ ")")::rev_res
     ) ctx2 (ctx1,[]) in
-  name_context, String.concat " " res
+  name_context, String.concat " " (List.rev rev_res)
 
 let rec pretty_print_term status ctxt =
   function
@@ -837,8 +962,8 @@ let rec pretty_print_term status ctxt =
       if pl = [] then
        "error \"Case analysis over empty type\""
       else
-       "case " ^ pretty_print_term status ctxt matched ^ " of\n" ^
-         String.concat "\n"
+       "case " ^ pretty_print_term status ctxt matched ^ " of {\n" ^
+         String.concat " ;\n"
            (HExtlib.list_mapi
              (fun (bound_names,rhs) i ->
                let ref = NReference.mk_constructor (i+1) r in
@@ -849,13 +974,15 @@ let rec pretty_print_term status ctxt =
                 pretty_print_term status ctxt rhs
                in
                  "  " ^ name ^ " " ^ bound_names ^ " -> " ^ body
-             ) pl)
-    | Skip t
-    | TLambda t -> pretty_print_term status (""@::ctxt) t
+             ) pl) ^ "}\n  "
+    | Skip t -> pretty_print_term status ("[[skipped]]"@::ctxt) t
+    | TLambda (name,t) ->
+       let name = capitalize `TypeVariable name in
+        pretty_print_term status (name@::ctxt) t
     | Inst t -> pretty_print_term status ctxt t
 ;;
 
-let rec pp_obj status (ref,obj_kind) =
+let rec pp_obj status (_,ref,obj_kind) =
   let pretty_print_context ctx =
     String.concat " " (List.rev (snd
       (List.fold_right
@@ -899,7 +1026,7 @@ let rec pp_obj status (ref,obj_kind) =
  | Algebraic il ->
     String.concat "\n"
      (List.map
-      (fun ref,left,right,cl ->
+      (fun _,ref,left,right,cl ->
         "data " ^ pp_ref status ref ^ " " ^
         pretty_print_context (right@left) ^ " where\n  " ^
         String.concat "\n  " (List.map
@@ -907,15 +1034,47 @@ let rec pp_obj status (ref,obj_kind) =
            let namectx = namectx_of_ctx left in
             pp_ref status ref ^ " :: " ^
              pretty_print_type status namectx tys
-         ) cl
-      )) il)
+         ) cl) ^ "\n    deriving (Prelude.Show)"
+      ) il)
  (* inductive and records missing *)
 
+let rec infos_of (info,_,obj_kind) =
+ info @
+  match obj_kind with
+     LetRec l -> List.concat (List.map (fun (infos,_,_) -> infos) l)
+   | Algebraic l -> List.concat (List.map (fun (infos,_,_,_,_) -> infos) l)
+   | _ -> []
+
 let haskell_of_obj status (uri,_,_,_,_ as obj) =
  let status, obj = object_of status obj in
   status,
    match obj with
-      Erased -> "-- [?] " ^ NUri.name_of_uri uri ^ " erased due to term being propositionally irrelevant.\n"
-    | OutsideTheory -> "-- [?] " ^ NUri.name_of_uri uri ^ " erased due to image of term under extraction residing outside Fω.\n"
-    | Failure msg -> "-- [?] " ^ NUri.name_of_uri uri ^ " FAILURE: " ^ msg ^ "\n"
-    | Success o -> pp_obj status o ^ "\n"
+      Erased -> "-- [?] " ^ NUri.name_of_uri uri ^ " erased due to term being propositionally irrelevant.\n",[]
+    | OutsideTheory -> "-- [?] " ^ NUri.name_of_uri uri ^ " erased due to image of term under extraction residing outside Fω.\n",[]
+    | Failure msg -> "-- [?] " ^ NUri.name_of_uri uri ^ " FAILURE: " ^ msg ^ "\n",[]
+    | Success o -> pp_obj status o ^ "\n", infos_of o
+
+let refresh_uri_in_typ ~refresh_uri_in_reference =
+ let rec refresh_uri_in_typ =
+  function
+   | Var _
+   | Unit
+   | Top as t -> t
+   | TConst r -> TConst (refresh_uri_in_reference r)
+   | Arrow (t1,t2) -> Arrow (refresh_uri_in_typ t1, refresh_uri_in_typ t2)
+   | TSkip t -> TSkip (refresh_uri_in_typ t)
+   | Forall (n,k,t) -> Forall (n,k,refresh_uri_in_typ t)
+   | TAppl tl -> TAppl (List.map refresh_uri_in_typ tl)
+ in
+  refresh_uri_in_typ
+
+let refresh_uri_in_info ~refresh_uri_in_reference infos =
+ List.map
+  (function (ref,(ctx,typ)) ->
+    let typ =
+     match typ with
+        None -> None
+      | Some t -> Some (refresh_uri_in_typ ~refresh_uri_in_reference t)
+    in
+     refresh_uri_in_reference ref,(ctx,typ))
+  infos