]> matita.cs.unibo.it Git - helm.git/blob - matita/components/ng_kernel/nCicExtraction.ml
d68ff2c4408582c2610dc23d30471e64a0020895
[helm.git] / matita / components / ng_kernel / nCicExtraction.ml
1 (* Copyright (C) 2000, HELM Team.
2  * 
3  * This file is part of HELM, an Hypertextual, Electronic
4  * Library of Mathematics, developed at the Computer Science
5  * Department, University of Bologna, Italy.
6  * 
7  * HELM is free software; you can redistribute it and/or
8  * modify it under the terms of the GNU General Public License
9  * as published by the Free Software Foundation; either version 2
10  * of the License, or (at your option) any later version.
11  * 
12  * HELM is distributed in the hope that it will be useful,
13  * but WITHOUT ANY WARRANTY; without even the implied warranty of
14  * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the
15  * GNU General Public License for more details.
16  *
17  * You should have received a copy of the GNU General Public License
18  * along with HELM; if not, write to the Free Software
19  * Foundation, Inc., 59 Temple Place - Suite 330, Boston,
20  * MA  02111-1307, USA.
21  * 
22  * For details, see the HELM World-Wide-Web page,
23  * http://cs.unibo.it/helm/.
24  *)
25
26 (* $Id: cicPp.ml 7413 2007-05-29 15:30:53Z tassi $ *)
27
28 let fix_sorts t = t;;
29 let apply_subst subst t = assert (subst=[]); t;;
30
31 type typformerreference = NReference.reference
32 type reference = NReference.reference
33
34 type kind =
35   | Type
36   | KArrow of kind * kind
37   | KSkip of kind (* dropped abstraction *)
38
39 let rec size_of_kind =
40   function
41     | Type -> 1
42     | KArrow (l, r) -> 1 + size_of_kind l + size_of_kind r
43     | KSkip k -> size_of_kind k
44 ;;
45
46 let bracket size_of pp o =
47   if size_of o > 1 then
48     "(" ^ pp o ^ ")"
49   else
50     pp o
51 ;;
52
53 let rec pretty_print_kind =
54   function
55     | Type -> "*"
56     | KArrow (l, r) -> bracket size_of_kind pretty_print_kind l ^ " -> " ^ pretty_print_kind r
57     | KSkip k -> pretty_print_kind k
58 ;;
59
60 type typ =
61   | Var of int
62   | Unit
63   | Top
64   | TConst of typformerreference
65   | Arrow of typ * typ
66   | Skip of typ
67   | Forall of string * kind * typ
68   | TAppl of typ list
69
70 let rec size_of_type =
71   function
72     | Var v -> 1
73     | Unit -> 1
74     | Top -> 1
75     | TConst c -> 1
76     | Arrow _ -> 2
77     | Skip t -> size_of_type t
78     | Forall _ -> 2
79     | TAppl l -> 1
80 ;;
81
82 type term =
83   | Rel of int
84   | UnitTerm
85   | Const of reference
86   | Lambda of string * (* typ **) term
87   | Appl of term list
88   | LetIn of string * (* typ **) term * term
89   | Match of reference * term * term list
90   | TLambda of (* string **) term
91   | Inst of (*typ_former **) term
92 ;;
93
94 let rec size_of_term =
95   function
96     | Rel r -> 1
97     | UnitTerm -> 1
98     | Const c -> 1
99     | Lambda (name, body) -> 1 + size_of_term body
100     | Appl l -> List.length l
101     | LetIn (name, def, body) -> 1 + size_of_term def + size_of_term body
102     | Match (name, case, pats) -> 1 + size_of_term case + List.length pats
103     | TLambda t -> size_of_term t
104     | Inst t -> size_of_term t
105 ;;
106 let unitty =
107  NCic.Const (NReference.reference_of_spec (NUri.uri_of_string "cic:/matita/basics/types/unit.ind") (NReference.Ind (true,0,0)));;
108
109 (* None = dropped abstraction *)
110 type typ_context = (string * kind) option list
111 type term_context = (string * [`OfKind of kind | `OfType of typ]) option list
112
113 type typ_former_decl = typ_context * kind
114 type typ_former_def = typ_former_decl * typ
115
116 type term_former_decl = term_context * typ
117 type term_former_def = term_former_decl * term
118
119 type obj_kind =
120    TypeDeclaration of typ_former_decl
121  | TypeDefinition of typ_former_def
122  | TermDeclaration of term_former_decl
123  | TermDefinition of term_former_def
124  | LetRec of obj_kind list
125    (* name, left, right, constructors *)
126  | Algebraic of (string * typ_context * typ_context * (string * typ) list) list
127
128 type obj = NUri.uri * obj_kind
129
130 let rec classify_not_term status context t =
131  match NCicReduction.whd status ~subst:[] context t with
132   | NCic.Sort s ->
133      (match s with
134          NCic.Prop
135        | NCic.Type [`CProp,_] -> `PropKind
136        | NCic.Type [`Type,_] -> `Kind
137        | NCic.Type _ -> assert false)
138   | NCic.Prod (b,s,t) ->
139      (*CSC: using invariant on "_" *)
140      classify_not_term status ((b,NCic.Decl s)::context) t
141   | NCic.Implicit _
142   | NCic.LetIn _
143   | NCic.Lambda _
144   | NCic.Const (NReference.Ref (_,NReference.CoFix _))
145   | NCic.Appl [] -> assert false (* NOT POSSIBLE *)
146   | NCic.Match _
147   | NCic.Const (NReference.Ref (_,NReference.Fix _)) ->
148      (* be aware: we can be the head of an application *)
149      assert false (* TODO *)
150   | NCic.Meta _ -> assert false (* TODO *)
151   | NCic.Appl (he::_) -> classify_not_term status context he
152   | NCic.Rel n ->
153      let rec find_sort typ =
154       match NCicReduction.whd status ~subst:[] context (NCicSubstitution.lift status n typ) with
155         NCic.Sort NCic.Prop -> `Proposition
156       | NCic.Sort (NCic.Type [`CProp,_]) -> `Proposition
157       | NCic.Sort (NCic.Type [`Type,_]) ->
158          (*CSC: we could be more precise distinguishing the user provided
159                 minimal elements of the hierarchies and classify these
160                 as `Kind *)
161          `KindOrType
162       | NCic.Sort (NCic.Type _) -> assert false (* ALGEBRAIC *)
163       | NCic.Prod (_,_,t) ->
164          (* we skipped arguments of applications, so here we need to skip
165             products *)
166          find_sort t
167       | _ -> assert false (* NOT A SORT *)
168      in
169       (match List.nth context (n-1) with
170           _,NCic.Decl typ -> find_sort typ
171         | _,NCic.Def _ -> assert false (* IMPOSSIBLE *))
172   | NCic.Const (NReference.Ref (_,NReference.Decl) as ref) ->
173      let _,_,ty,_,_ = NCicEnvironment.get_checked_decl status ref in
174       (match classify_not_term status [] ty with
175         | `Proposition
176         | `Type -> assert false (* IMPOSSIBLE *)
177         | `Kind
178         | `KindOrType -> `Type
179         | `PropKind -> `Proposition)
180   | NCic.Const (NReference.Ref (_,NReference.Ind _) as ref) ->
181      let _,_,ityl,_,i = NCicEnvironment.get_checked_indtys status ref in
182      let _,_,arity,_ = List.nth ityl i in
183       (match classify_not_term status [] arity with
184         | `Proposition
185         | `Type
186         | `KindOrType -> assert false (* IMPOSSIBLE *)
187         | `Kind -> `Type
188         | `PropKind -> `Proposition)
189   | NCic.Const (NReference.Ref (_,NReference.Con _))
190   | NCic.Const (NReference.Ref (_,NReference.Def _)) ->
191      assert false (* IMPOSSIBLE *)
192 ;;
193
194 type not_term = [`Kind | `KindOrType | `PropKind | `Proposition | `Type];;
195
196 let classify status ~metasenv context t =
197  match NCicTypeChecker.typeof status ~metasenv ~subst:[] context t with
198   | NCic.Sort _ ->
199      (classify_not_term status context t : not_term :> [> not_term])
200   | ty ->
201      let ty = fix_sorts ty in
202       `Term
203         (match classify_not_term status context ty with
204           | `Proposition -> `Proof
205           | `Type -> `Term
206           | `KindOrType -> `TypeFormerOrTerm
207           | `Kind -> `TypeFormer
208           | `PropKind -> `PropFormer)
209 ;;
210       
211
212 let rec kind_of status ~metasenv context k =
213  match NCicReduction.whd status ~subst:[] context k with
214   | NCic.Sort NCic.Type _ -> Type
215   | NCic.Sort _ -> assert false (* NOT A KIND *)
216   | NCic.Prod (b,s,t) ->
217      (match classify status ~metasenv context s with
218        | `Kind ->
219            KArrow (kind_of status ~metasenv context s,
220             kind_of ~metasenv status ((b,NCic.Decl s)::context) t)
221        | `Type
222        | `KindOrType
223        | `Proposition
224        | `PropKind ->
225            KSkip (kind_of status ~metasenv ((b,NCic.Decl s)::context) t)
226        | `Term _ -> assert false (* IMPOSSIBLE *))
227   | NCic.Implicit _
228   | NCic.LetIn _ -> assert false (* IMPOSSIBLE *)
229   | NCic.Lambda _
230   | NCic.Rel _
231   | NCic.Const _ -> assert false (* NOT A KIND *)
232   | NCic.Appl _ -> assert false (* TODO: when head is a match/let rec;
233                                    otherwise NOT A KIND *)
234   | NCic.Meta _
235   | NCic.Match (_,_,_,_) -> assert false (* TODO *)
236 ;;
237
238 let rec skip_args status ~metasenv context =
239  function
240   | _,[] -> []
241   | [],_ -> assert false (* IMPOSSIBLE *)
242   | None::tl1,_::tl2 -> skip_args status ~metasenv context (tl1,tl2)
243   | _::tl1,arg::tl2 ->
244      match classify status ~metasenv context arg with
245       | `KindOrType
246       | `Type
247       | `Term `TypeFormer ->
248          Some arg::skip_args status ~metasenv context (tl1,tl2)
249       | `Kind
250       | `Proposition
251       | `PropKind -> None::skip_args status ~metasenv context (tl1,tl2)
252       | `Term _ -> assert false (* IMPOSSIBLE *)
253 ;;
254
255 module ReferenceMap = Map.Make(struct type t = NReference.reference let compare = NReference.compare end)
256
257 type db = (typ_context * typ option) ReferenceMap.t
258
259 class type g_status =
260  object
261   method extraction_db: db
262  end
263
264 class virtual status =
265  object
266   inherit NCic.status
267   val extraction_db = ReferenceMap.empty
268   method extraction_db = extraction_db
269   method set_extraction_db v = {< extraction_db = v >}
270   method set_extraction_status
271    : 'status. #g_status as 'status -> 'self
272    = fun o -> {< extraction_db = o#extraction_db >}
273  end
274
275 let rec split_kind_prods context =
276  function
277   | KArrow (so,ta)-> split_kind_prods (Some ("_",so)::context) ta
278   | KSkip ta -> split_kind_prods (None::context) ta
279   | Type -> context,Type
280 ;;
281
282 let rec split_typ_prods context =
283  function
284   | Arrow (so,ta)-> split_typ_prods (Some ("_",`OfType so)::context) ta
285   | Forall (name,so,ta)-> split_typ_prods (Some (name,`OfKind so)::context) ta
286   | Skip ta -> split_typ_prods (None::context) ta
287   | _ as t -> context,t
288 ;;
289
290 let rec glue_ctx_typ ctx typ =
291  match ctx with
292   | [] -> typ
293   | Some (_,`OfType so)::ctx -> glue_ctx_typ ctx (Arrow (so,typ))
294   | Some (name,`OfKind so)::ctx -> glue_ctx_typ ctx (Forall (name,so,typ))
295   | None::ctx -> glue_ctx_typ ctx (Skip typ)
296 ;;
297
298 let rec split_typ_lambdas status n ~metasenv context typ =
299  if n = 0 then context,typ
300  else
301   match NCicReduction.whd status ~delta:max_int ~subst:[] context typ with
302    | NCic.Lambda (name,s,t) ->
303       split_typ_lambdas status (n-1) ~metasenv ((name,NCic.Decl s)::context) t
304    | t ->
305       (* eta-expansion required *)
306       let ty = NCicTypeChecker.typeof status ~metasenv ~subst:[] context t in
307       match NCicReduction.whd status ~delta:max_int ~subst:[] context ty with
308        | NCic.Prod (name,typ,_) ->
309           split_typ_lambdas status (n-1) ~metasenv
310            ((name,NCic.Decl typ)::context)
311            (NCicUntrusted.mk_appl t [NCic.Rel 1])
312        | _ -> assert false (* IMPOSSIBLE *)
313 ;;
314
315
316 let context_of_typformer status ~metasenv context =
317  function
318     NCic.Const (NReference.Ref (_,NReference.Ind _) as ref)
319   | NCic.Const (NReference.Ref (_,NReference.Def _) as ref)
320   | NCic.Const (NReference.Ref (_,NReference.Decl) as ref)
321   | NCic.Const (NReference.Ref (_,NReference.Fix _) as ref) ->
322      (try fst (ReferenceMap.find ref status#extraction_db)
323       with
324        Not_found -> assert false (* IMPOSSIBLE *))
325   | NCic.Match _ -> assert false (* TODO ???? *)
326   | NCic.Rel n ->
327      let typ =
328       match List.nth context (n-1) with
329          _,NCic.Decl typ -> typ
330        | _,NCic.Def _ -> assert false (* IMPOSSIBLE *) in
331      let typ_ctx = snd (HExtlib.split_nth n context) in
332      let typ = kind_of status ~metasenv typ_ctx typ in
333       fst (split_kind_prods [] typ)
334   | NCic.Meta _ -> assert false (* TODO *)
335   | NCic.Const (NReference.Ref (_,NReference.Con _))
336   | NCic.Const (NReference.Ref (_,NReference.CoFix _))
337   | NCic.Sort _ | NCic.Implicit _ | NCic.Lambda _ | NCic.LetIn _
338   | NCic.Appl _ | NCic.Prod _ ->
339      assert false (* IMPOSSIBLE *)
340
341 let rec typ_of status ~metasenv context k =
342  match NCicReduction.whd status ~delta:max_int ~subst:[] context k with
343   | NCic.Prod (b,s,t) ->
344      (* CSC: non-invariant assumed here about "_" *)
345      (match classify status ~metasenv context s with
346        | `Kind ->
347            Forall (b, kind_of status ~metasenv context s,
348             typ_of ~metasenv status ((b,NCic.Decl s)::context) t)
349        | `Type
350        | `KindOrType -> (* ??? *)
351            Arrow (typ_of status ~metasenv context s,
352             typ_of status ~metasenv ((b,NCic.Decl s)::context) t)
353        | `PropKind
354        | `Proposition ->
355            Skip (typ_of status ~metasenv ((b,NCic.Decl s)::context) t)
356        | `Term _ -> assert false (* IMPOSSIBLE *))
357   | NCic.Sort _
358   | NCic.Implicit _
359   | NCic.LetIn _ -> assert false (* IMPOSSIBLE *)
360   | NCic.Lambda _ -> assert false (* LAMBDA-LIFT INNER DECLARATION *)
361   | NCic.Rel n -> Var n
362   | NCic.Const ref -> TConst ref
363   | NCic.Appl (he::args) ->
364      let he_context = context_of_typformer status ~metasenv context he in
365      TAppl (typ_of status ~metasenv context he ::
366       List.map
367        (function None -> Unit | Some ty -> typ_of status ~metasenv context ty)
368        (skip_args status ~metasenv context (List.rev he_context,args)))
369   | NCic.Appl _ -> assert false (* TODO: when head is a match/let rec;
370                                    otherwise NOT A TYPE *)
371   | NCic.Meta _
372   | NCic.Match (_,_,_,_) -> assert false (* TODO *)
373 ;;
374
375 let rec fomega_subst k t1 =
376  function
377   | Var n ->
378      if k=n then t1
379      else if n < k then Var n
380      else Var (n-1)
381   | Top -> Top
382   | TConst ref -> TConst ref
383   | Unit -> Unit
384   | Arrow (ty1,ty2) -> Arrow (fomega_subst k t1 ty1, fomega_subst (k+1) t1 ty2)
385   | Skip t -> Skip (fomega_subst (k+1) t1 t)
386   | Forall (n,kind,t) -> Forall (n,kind,fomega_subst (k+1) t1 t)
387   | TAppl args -> TAppl (List.map (fomega_subst k t1) args)
388
389 let fomega_lookup status ref = snd (ReferenceMap.find ref status#extraction_db)
390
391 let rec fomega_whd status ty =
392  match ty with
393  | TConst r ->
394     (match fomega_lookup status r with
395         None -> ty
396       | Some ty -> fomega_whd status ty)
397  | TAppl (TConst r::args) ->
398     (match fomega_lookup status r with
399         None -> ty
400       | Some ty -> fomega_whd status (List.fold_right (fomega_subst 1) args ty))
401  | _ -> ty
402
403 let rec term_of status ~metasenv context =
404  function
405   | NCic.Implicit _
406   | NCic.Sort _
407   | NCic.Prod _ -> assert false (* IMPOSSIBLE *)
408   | NCic.Lambda (b,ty,bo) ->
409      (* CSC: non-invariant assumed here about "_" *)
410      (match classify status ~metasenv context ty with
411        | `Kind ->
412            TLambda (term_of status ~metasenv ((b,NCic.Decl ty)::context) bo)
413        | `KindOrType (* ??? *)
414        | `Type ->
415            Lambda (b, term_of status ~metasenv ((b,NCic.Decl ty)::context) bo)
416        | `PropKind
417        | `Proposition ->
418            (* CSC: LAZY ??? *)
419            term_of status ~metasenv ((b,NCic.Decl ty)::context) bo
420        | `Term _ -> assert false (* IMPOSSIBLE *))
421   | NCic.LetIn (b,ty,t,bo) ->
422      (match classify status ~metasenv context t with
423        | `Term `TypeFormerOrTerm (* ???? *)
424        | `Term `Term ->
425           LetIn (b,term_of status ~metasenv context t,
426            term_of status ~metasenv ((b,NCic.Def (t,ty))::context) bo)
427        | `Kind
428        | `Type
429        | `KindOrType
430        | `PropKind
431        | `Proposition
432        | `Term `PropFormer
433        | `Term `TypeFormer
434        | `Term `Proof ->
435          (* not in programming languages, we expand it *)
436          term_of status ~metasenv context
437           (NCicSubstitution.subst status ~avoid_beta_redexes:true t bo))
438   | NCic.Rel n -> Rel n
439   | NCic.Const ref -> Const ref
440   | NCic.Appl (he::args) ->
441      eat_args status metasenv
442       (term_of status ~metasenv context he) context
443       (typ_of status ~metasenv context
444         (NCicTypeChecker.typeof status ~metasenv ~subst:[] context he))
445       args
446 (*
447      let he_context = context_of_typformer status ~metasenv context he in
448      let process_args he =
449       function
450          `Empty -> he
451        | `Inst tl -> Inst (process_args he tl)
452        | `Appl (arg,tl) -> process_args (Appl (he,... arg)) tl
453      in
454      Appl (typ_of status ~metasenv context he ::
455       process_args (typ_of status ~metasenv context he)
456        (skip_term_args status ~metasenv context (List.rev he_context,args))
457 *)
458   | NCic.Appl _ -> assert false (* TODO: when head is a match/let rec;
459                                    otherwise NOT A TYPE *)
460   | NCic.Meta _ -> assert false (* TODO *)
461   | NCic.Match (ref,_,t,pl) ->
462      Match (ref,term_of status ~metasenv context t,
463       List.map (term_of status ~metasenv context) pl)
464 and eat_args status metasenv acc context tyhe =
465  function
466     [] -> acc
467   | arg::tl ->
468      let mk_appl f x =
469       match f with
470          Appl args -> Appl (args@[x])
471        | _ -> Appl [f;x]
472      in
473      match fomega_whd status tyhe with
474         Arrow (s,t) ->
475          let arg =
476           match s with
477              Unit -> UnitTerm
478            | _ -> term_of status ~metasenv context arg in
479          eat_args status metasenv (mk_appl acc arg) context t tl
480       | Forall (_,_,t) ->
481          eat_args status metasenv (Inst acc)
482           context (fomega_subst 1 (typ_of status ~metasenv context arg) t) tl
483       | Skip t ->
484          eat_args status metasenv acc context t tl
485       | Top -> assert false (*TODO: HOW??*)
486       | Unit | Var _ | TConst _ | TAppl _ -> assert false (* NOT A PRODUCT *)
487 ;;
488
489 type 'a result =
490   | Erased
491   | OutsideTheory
492   | Failure of string
493   | Success of 'a
494 ;;
495
496 let object_of_constant status ~metasenv uri height bo ty =
497   match classify status ~metasenv [] ty with
498     | `Kind ->
499         let ty = kind_of status ~metasenv [] ty in
500         let ctx0,res = split_kind_prods [] ty in
501         let ctx,bo =
502          split_typ_lambdas status ~metasenv (List.length ctx0) [] bo in
503         (match classify status ~metasenv ctx bo with
504           | `Type
505           | `KindOrType -> (* ?? no kind formers in System F_omega *)
506               let nicectx =
507                List.map2
508                 (fun p1 n ->
509                   HExtlib.map_option (fun (_,k) ->
510                    (*CSC: BUG here, clashes*)
511                    String.uncapitalize (fst n),k) p1)
512                 ctx0 ctx
513               in
514               (* BUG here: for mutual type definitions the spec is not good *)
515               let ref =
516                NReference.reference_of_spec uri (NReference.Def height) in
517               let bo = typ_of status ~metasenv ctx bo in
518                status#set_extraction_db
519                 (ReferenceMap.add ref (nicectx,Some bo)
520                   status#extraction_db),
521                Success (uri,TypeDefinition((nicectx,res),bo))
522           | `Kind -> status, Erased (* DPM: but not really, more a failure! *)
523           | `PropKind
524           | `Proposition -> status, Erased
525           | `Term _ -> status, Failure "Body of type lambda classified as a term.  This is a bug.")
526     | `PropKind
527     | `Proposition -> status, Erased
528     | `KindOrType (* ??? *)
529     | `Type ->
530        (* CSC: TO BE FINISHED, REF NON REGISTERED *)
531        let ty = typ_of status ~metasenv [] ty in
532         status,
533         Success (uri, TermDefinition (split_typ_prods [] ty, term_of status ~metasenv [] bo))
534     | `Term _ -> status, Failure "Non-term classified as a term.  This is a bug."
535 ;;
536
537 let object_of_inductive status ~metasenv uri ind leftno il =
538  let status,_,rev_tyl =
539   List.fold_left
540    (fun (status,i,res) (_,name,arity,cl) ->
541      match classify_not_term status [] arity with
542       | `Proposition
543       | `KindOrType
544       | `Type -> assert false (* IMPOSSIBLE *)
545       | `PropKind -> status,i+1,res
546       | `Kind ->
547           let arity = kind_of status ~metasenv [] arity in
548           let ctx,_ = split_kind_prods [] arity in
549           let right,left = HExtlib.split_nth (List.length ctx - leftno) ctx in
550           let ref =
551            NReference.reference_of_spec uri (NReference.Ind (ind,i,leftno)) in
552           let status =
553            status#set_extraction_db
554             (ReferenceMap.add ref (ctx,None) status#extraction_db) in
555           let cl =
556            List.map
557             (fun (_,name,ty) ->
558               let ctx,ty =
559                NCicReduction.split_prods status ~subst:[] [] leftno ty in
560               let ty = typ_of status ~metasenv ctx ty in
561                name,ty
562             ) cl
563           in
564            status,i+1,(name,left,right,cl)::res
565    ) (status,0,[]) il
566  in
567   match rev_tyl with
568      [] -> status,Erased
569    | _ -> status, Success (uri, Algebraic (List.rev rev_tyl))
570 ;;
571
572 let object_of status (uri,height,metasenv,subst,obj_kind) =
573   let obj_kind = apply_subst subst obj_kind in
574       match obj_kind with
575         | NCic.Constant (_,_,None,ty,_) ->
576           (match classify status ~metasenv [] ty with
577             | `Kind ->
578               let ty = kind_of status ~metasenv [] ty in
579               let ctx,_ as res = split_kind_prods [] ty in
580               let ref = NReference.reference_of_spec uri NReference.Decl in
581                 status#set_extraction_db
582                   (ReferenceMap.add ref (ctx,None) status#extraction_db), Success (uri, TypeDeclaration res)
583             | `PropKind
584             | `Proposition -> status, Erased
585             | `Type
586             | `KindOrType (*???*) ->
587               let ty = typ_of status ~metasenv [] ty in
588                 status, Success (uri, TermDeclaration (split_typ_prods [] ty))
589             | `Term _ -> status, Failure "Type classified as a term.  This is a bug.")
590         | NCic.Constant (_,_,Some bo,ty,_) ->
591           object_of_constant status ~metasenv uri height bo ty
592         | NCic.Fixpoint (_fix_or_cofix,fs,_) ->
593           let status,objs =
594             List.fold_right
595             (fun (_,_name,_,ty,bo) (status,res) ->
596               let status,obj = object_of_constant ~metasenv status uri height bo ty in
597                 match obj with
598                   | Success (_uri,obj) -> status, obj::res
599                   | _ -> status, res) fs (status,[])
600           in
601             status, Success (uri,LetRec objs)
602         | NCic.Inductive (ind,leftno,il,_) ->
603            object_of_inductive status ~metasenv uri ind leftno il
604
605 (************************ HASKELL *************************)
606
607 (* -----------------------------------------------------------------------------
608  * Helper functions I can't seem to find anywhere in the OCaml stdlib?
609  * -----------------------------------------------------------------------------
610  *)
611 let (|>) f g =
612   fun x -> g (f x)
613 ;;
614
615 let curry f x y =
616   f (x, y)
617 ;;
618
619 let uncurry f (x, y) =
620   f x y
621 ;;
622
623 let rec char_list_of_string s =
624   let l = String.length s in
625   let rec aux buffer s =
626     function
627       | 0 -> buffer
628       | m -> aux (s.[m - 1]::buffer) s (m - 1)
629   in
630     aux [] s l
631 ;;
632
633 let string_of_char_list s =
634   let rec aux buffer =
635     function
636       | []    -> buffer
637       | x::xs -> aux (String.make 1 x ^ buffer) xs
638   in
639     aux "" s
640 ;;
641
642 (* ----------------------------------------------------------------------------
643  * Haskell name management: prettyfying valid and idiomatic Haskell identifiers
644  * and type variable names.
645  * ----------------------------------------------------------------------------
646  *)
647
648 let remove_underscores_and_mark =
649   let rec aux char_list_buffer positions_buffer position =
650     function
651       | []    -> (string_of_char_list char_list_buffer, positions_buffer)
652       | x::xs ->
653         if x == '_' then
654           aux char_list_buffer (position::positions_buffer) position xs
655         else
656           aux (x::char_list_buffer) positions_buffer (position + 1) xs
657   in
658     aux [] [] 0
659 ;;
660
661 let rec capitalize_marked_positions s =
662   function
663     | []    -> s
664     | x::xs ->
665       if x < String.length s then
666         let c = Char.uppercase (String.get s x) in
667         let _ = String.set s x c in
668           capitalize_marked_positions s xs
669       else
670         capitalize_marked_positions s xs
671 ;;
672
673 let contract_underscores_and_capitalise =
674   char_list_of_string |>
675   remove_underscores_and_mark |>
676   uncurry capitalize_marked_positions
677 ;;
678
679 let idiomatic_haskell_type_name_of_string =
680   contract_underscores_and_capitalise |>
681   String.capitalize
682 ;;
683
684 let idiomatic_haskell_term_name_of_string =
685   contract_underscores_and_capitalise |>
686   String.uncapitalize
687 ;;
688
689 (*CSC: code to be changed soon when we implement constructors and
690   we fix the code for term application *)
691 let classify_reference status ref =
692   if ReferenceMap.mem ref status#extraction_db then
693     `TypeName
694   else
695     `FunctionName
696 ;;
697
698 let capitalize classification name =
699   match classification with
700     | `Constructor
701     | `TypeName -> idiomatic_haskell_type_name_of_string name
702     | `FunctionName -> idiomatic_haskell_term_name_of_string name
703 ;;
704
705 let pp_ref status ref =
706  capitalize (classify_reference status ref)
707   (NCicPp.r2s status false ref)
708
709 let name_of_uri classification uri =
710  capitalize classification (NUri.name_of_uri uri)
711
712 (* cons avoid duplicates *)
713 let rec (@:::) name l =
714  if name <> "" (* propositional things *) && name.[0] = '_' then
715   let name = String.sub name 1 (String.length name - 1) in
716   let name = if name = "" then "a" else name in
717    name @::: l
718  else if List.mem name l then (name ^ "'") @::: l
719  else if name="" then "[erased]",l else name,l
720 ;;
721
722 let (@::) x l = let x,l = x @::: l in x::l;;
723
724 let rec pretty_print_type status ctxt =
725   function
726     | Var n -> List.nth ctxt (n-1)
727     | Unit -> "()"
728     | Top -> assert false (* ??? *)
729     | TConst ref -> pp_ref status ref
730     | Arrow (t1,t2) ->
731         bracket size_of_type (pretty_print_type status ctxt) t1 ^ " -> " ^
732          pretty_print_type status ("_"::ctxt) t2
733     | Skip t -> pretty_print_type status ("_"::ctxt) t
734     | Forall (name, kind, t) ->
735       (*CSC: BUG HERE: avoid clashes due to uncapitalisation*)
736       let name = String.uncapitalize name in
737         if size_of_kind kind > 1 then
738           "forall (" ^ name ^ " :: " ^ pretty_print_kind kind ^ "). " ^ pretty_print_type status (name@::ctxt) t
739         else
740           "forall " ^ name ^ ". " ^ pretty_print_type status (name@::ctxt) t
741    | TAppl tl -> String.concat " " (List.map (pretty_print_type status ctxt) tl)
742
743 let rec pretty_print_term status ctxt =
744   function
745     | Rel n -> List.nth ctxt (n-1)
746     | UnitTerm -> "()"
747     | Const ref -> pp_ref status ref
748     | Lambda (name,t) -> "\\" ^ name ^ " -> " ^ pretty_print_term status (name@::ctxt) t
749     | Appl tl -> String.concat " " (List.map (bracket size_of_term (pretty_print_term status ctxt)) tl)
750     | LetIn (name,s,t) ->
751       "let " ^ name ^ " = " ^ pretty_print_term status ctxt s ^ " in " ^ pretty_print_term status (name@::ctxt) t
752     | Match (r,matched,pl) ->
753       if pl = [] then
754        "error \"Case analysis over empty type\""
755       else
756       let constructors, leftno =
757       let _,leftno,tys,_,n = NCicEnvironment.get_checked_indtys status r in
758       let _,_,_,cl  = List.nth tys n in
759         cl,leftno
760       in
761         let rec eat_branch n ty pat =
762           match (ty, pat) with
763           | NCic.Prod (_, _, t), _ when n > 0 -> eat_branch (pred n) t pat 
764           | NCic.Prod (_, _, t), Lambda (name, t') ->
765             (*CSC: BUG HERE; WHAT IF SOME ARGUMENTS ARE DELETED?*)
766             let cv, rhs = eat_branch 0 t t' in
767               name :: cv, rhs
768           | _, _ -> [], pat
769         in
770           let j = ref 0 in
771           let patterns =
772             try
773               List.map2
774                 (fun (_, name, ty) pat -> incr j; name, eat_branch leftno ty pat) constructors pl
775             with Invalid_argument _ -> assert false
776           in
777             "case " ^ pretty_print_term status ctxt matched ^ " of\n" ^
778               String.concat "\n"
779                 (List.map
780                   (fun (name,(bound_names,rhs)) ->
781                     let pattern,body =
782                     (*CSC: BUG avoid name clashes *)
783                     String.concat " " (String.capitalize name::bound_names),
784                       pretty_print_term status ((List.rev bound_names)@ctxt) rhs
785                     in
786                       "  " ^ pattern ^ " -> " ^ body
787                   ) patterns)
788     | TLambda t -> pretty_print_term status ctxt t
789     | Inst t -> pretty_print_term status ctxt t
790 ;;
791
792 (*
793 type term_context = (string * [`OfKind of kind | `OfType of typ]) option list
794
795 type term_former_def = term_context * term * typ
796 type term_former_decl = term_context * typ
797 *)
798
799 let rec pp_obj status (uri,obj_kind) =
800   let pretty_print_context ctx =
801     String.concat " " (List.rev (snd
802       (List.fold_right
803        (fun (x,kind) (l,res) ->
804          let x,l = x @:::l in
805          if size_of_kind kind > 1 then
806           x::l,("(" ^ x ^ " :: " ^ pretty_print_kind kind ^ ")")::res
807          else
808           x::l,x::res)
809        (HExtlib.filter_map (fun x -> x) ctx) ([],[]))))
810   in
811  let namectx_of_ctx ctx =
812   List.fold_right (@::)
813    (List.map (function None -> "" | Some (x,_) -> x) ctx) [] in
814  match obj_kind with
815    TypeDeclaration (ctx,_) ->
816     (* data?? unsure semantics: inductive type without constructor, but
817        not matchable apparently *)
818     if List.length ctx = 0 then
819       "data " ^ name_of_uri `TypeName uri
820     else
821       "data " ^ name_of_uri `TypeName uri ^ " " ^ pretty_print_context ctx
822  | TypeDefinition ((ctx, _),ty) ->
823     let namectx = namectx_of_ctx ctx in
824     if List.length ctx = 0 then
825       "type " ^ name_of_uri `TypeName uri ^ " = " ^ pretty_print_type status namectx ty
826     else
827       "type " ^ name_of_uri `TypeName uri ^ " " ^ pretty_print_context ctx ^ " = " ^ pretty_print_type status namectx ty
828  | TermDeclaration (ctx,ty) ->
829     let name = name_of_uri `FunctionName uri in
830       name ^ " :: " ^ pretty_print_type status [] (glue_ctx_typ ctx ty) ^ "\n" ^
831       name ^ " = error \"The declaration `" ^ name ^ "' has yet to be defined.\""
832  | TermDefinition ((ctx,ty),bo) ->
833     let name = name_of_uri `FunctionName uri in
834     let namectx = namectx_of_ctx ctx in
835     (*CSC: BUG here *)
836     name ^ " :: " ^ pretty_print_type status namectx (glue_ctx_typ ctx ty) ^ "\n" ^
837     name ^ " = " ^ pretty_print_term status namectx bo
838  | LetRec l ->
839     (*CSC: BUG always uses the name of the URI *)
840     String.concat "\n" (List.map (fun obj -> pp_obj status (uri,obj)) l)
841  | Algebraic il ->
842     String.concat "\n"
843      (List.map
844       (fun _name,left,right,cl ->
845         (*CSC: BUG always uses the name of the URI *)
846         "data " ^ name_of_uri `TypeName uri ^ " " ^
847         pretty_print_context (right@left) ^ " where\n  " ^
848         String.concat "\n  " (List.map
849          (fun name,tys ->
850            let namectx = namectx_of_ctx left in
851             capitalize `Constructor name ^ " :: " ^
852              pretty_print_type status namectx tys
853          ) cl
854       )) il)
855  (* inductive and records missing *)
856
857 let haskell_of_obj status (uri,_,_,_,_ as obj) =
858  let status, obj = object_of status obj in
859   status,
860    match obj with
861       Erased -> "-- [?] " ^ NUri.name_of_uri uri ^ " erased due to term being propositionally irrelevant.\n"
862     | OutsideTheory -> "-- [?] " ^ NUri.name_of_uri uri ^ " erased due to image of term under extraction residing outside Fω.\n"
863     | Failure msg -> "-- [?] " ^ NUri.name_of_uri uri ^ " FAILURE: " ^ msg ^ "\n"
864     | Success o -> pp_obj status o ^ "\n"