1 (* Copyright (C) 2000, HELM Team.
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.
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.
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.
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,
22 * For details, see the HELM World-Wide-Web page,
23 * http://cs.unibo.it/helm/.
26 (* $Id: cicPp.ml 7413 2007-05-29 15:30:53Z tassi $ *)
29 let apply_subst subst t = assert (subst=[]); t;;
31 type typformerreference = NReference.reference
32 type reference = NReference.reference
36 | KArrow of kind * kind
37 | KSkip of kind (* dropped abstraction *)
39 let rec size_of_kind =
42 | KArrow (l, r) -> 1 + size_of_kind l + size_of_kind r
43 | KSkip k -> size_of_kind k
46 let bracket size_of pp o =
53 let rec pretty_print_kind =
56 | KArrow (l, r) -> bracket size_of_kind pretty_print_kind l ^ " -> " ^ pretty_print_kind r
57 | KSkip k -> pretty_print_kind k
64 | TConst of typformerreference
67 | Forall of string * kind * typ
70 let rec size_of_type =
77 | TSkip t -> size_of_type t
82 (* None = dropped abstraction *)
83 type typ_context = (string * kind) option list
84 type term_context = (string * [`OfKind of kind | `OfType of typ]) option list
86 type typ_former_decl = typ_context * kind
87 type typ_former_def = typ_former_decl * typ
93 | Lambda of string * (* typ **) term
95 | LetIn of string * (* typ **) term * term
96 | Match of reference * term * (term_context * term) list
98 | TLambda of string * term
99 | Inst of (*typ_former **) term
101 | UnsafeCoerce of term
104 type term_former_decl = term_context * typ
105 type term_former_def = term_former_decl * term
109 Appl args -> Appl (args@[x])
112 let rec size_of_term =
117 | Lambda (_, body) -> 1 + size_of_term body
118 | Appl l -> List.length l
119 | LetIn (_, def, body) -> 1 + size_of_term def + size_of_term body
120 | Match (_, case, pats) -> 1 + size_of_term case + List.length pats
122 | TLambda (_,t) -> size_of_term t
123 | Inst t -> size_of_term t
124 | Skip t -> size_of_term t
125 | UnsafeCoerce t -> 1 + size_of_term t
128 NCic.Const (NReference.reference_of_spec (NUri.uri_of_string "cic:/matita/basics/types/unit.ind") (NReference.Ind (true,0,0)));;
131 TypeDeclaration of typ_former_decl
132 | TypeDefinition of typ_former_def
133 | TermDeclaration of term_former_decl
134 | TermDefinition of term_former_def
135 | LetRec of (NReference.reference * obj_kind) list
136 (* reference, left, right, constructors *)
137 | Algebraic of (NReference.reference * typ_context * typ_context * (NReference.reference * typ) list) list
139 type obj = NReference.reference * obj_kind
141 (* For LetRec and Algebraic blocks *)
143 NReference.reference_of_string "cic:/matita/dummy/dummy.ind(1,1,1)"
145 let rec classify_not_term status context t =
146 match NCicReduction.whd status ~subst:[] context t with
150 | NCic.Type [`CProp,_] -> `PropKind
151 | NCic.Type [`Type,_] -> `Kind
152 | NCic.Type _ -> assert false)
153 | NCic.Prod (b,s,t) ->
154 (*CSC: using invariant on "_" *)
155 classify_not_term status ((b,NCic.Decl s)::context) t
159 | NCic.Const (NReference.Ref (_,NReference.CoFix _))
160 | NCic.Appl [] -> assert false (* NOT POSSIBLE *)
162 | NCic.Const (NReference.Ref (_,NReference.Fix _)) ->
163 (* be aware: we can be the head of an application *)
164 assert false (* TODO *)
165 | NCic.Meta _ -> assert false (* TODO *)
166 | NCic.Appl (he::_) -> classify_not_term status context he
168 let rec find_sort typ =
169 match NCicReduction.whd status ~subst:[] context (NCicSubstitution.lift status n typ) with
170 NCic.Sort NCic.Prop -> `Proposition
171 | NCic.Sort (NCic.Type [`CProp,_]) -> `Proposition
172 | NCic.Sort (NCic.Type [`Type,_]) ->
173 (*CSC: we could be more precise distinguishing the user provided
174 minimal elements of the hierarchies and classify these
177 | NCic.Sort (NCic.Type _) -> assert false (* ALGEBRAIC *)
178 | NCic.Prod (_,_,t) ->
179 (* we skipped arguments of applications, so here we need to skip
182 | _ -> assert false (* NOT A SORT *)
184 (match List.nth context (n-1) with
185 _,NCic.Decl typ -> find_sort typ
186 | _,NCic.Def _ -> assert false (* IMPOSSIBLE *))
187 | NCic.Const (NReference.Ref (_,NReference.Decl) as ref) ->
188 let _,_,ty,_,_ = NCicEnvironment.get_checked_decl status ref in
189 (match classify_not_term status [] ty with
191 | `Type -> assert false (* IMPOSSIBLE *)
193 | `KindOrType -> `Type
194 | `PropKind -> `Proposition)
195 | NCic.Const (NReference.Ref (_,NReference.Ind _) as ref) ->
196 let _,_,ityl,_,i = NCicEnvironment.get_checked_indtys status ref in
197 let _,_,arity,_ = List.nth ityl i in
198 (match classify_not_term status [] arity with
201 | `KindOrType -> assert false (* IMPOSSIBLE *)
203 | `PropKind -> `Proposition)
204 | NCic.Const (NReference.Ref (_,NReference.Con _))
205 | NCic.Const (NReference.Ref (_,NReference.Def _)) ->
206 assert false (* IMPOSSIBLE *)
209 type not_term = [`Kind | `KindOrType | `PropKind | `Proposition | `Type];;
211 let classify status ~metasenv context t =
212 match NCicTypeChecker.typeof status ~metasenv ~subst:[] context t with
214 (classify_not_term status context t : not_term :> [> not_term])
216 let ty = fix_sorts ty in
218 (match classify_not_term status context ty with
219 | `Proposition -> `Proof
221 | `KindOrType -> `TypeFormerOrTerm
222 | `Kind -> `TypeFormer
223 | `PropKind -> `PropFormer)
227 let rec kind_of status ~metasenv context k =
228 match NCicReduction.whd status ~subst:[] context k with
229 | NCic.Sort NCic.Type _ -> Type
230 | NCic.Sort _ -> assert false (* NOT A KIND *)
231 | NCic.Prod (b,s,t) ->
232 (match classify status ~metasenv context s with
234 KArrow (kind_of status ~metasenv context s,
235 kind_of ~metasenv status ((b,NCic.Decl s)::context) t)
240 KSkip (kind_of status ~metasenv ((b,NCic.Decl s)::context) t)
241 | `Term _ -> assert false (* IMPOSSIBLE *))
243 | NCic.LetIn _ -> assert false (* IMPOSSIBLE *)
246 | NCic.Const _ -> assert false (* NOT A KIND *)
247 | NCic.Appl _ -> assert false (* TODO: when head is a match/let rec;
248 otherwise NOT A KIND *)
250 | NCic.Match (_,_,_,_) -> assert false (* TODO *)
253 let rec skip_args status ~metasenv context =
256 | [],_ -> assert false (* IMPOSSIBLE *)
257 | None::tl1,_::tl2 -> skip_args status ~metasenv context (tl1,tl2)
259 match classify status ~metasenv context arg with
262 | `Term `TypeFormer ->
263 Some arg::skip_args status ~metasenv context (tl1,tl2)
266 | `PropKind -> None::skip_args status ~metasenv context (tl1,tl2)
267 | `Term _ -> assert false (* IMPOSSIBLE *)
270 module ReferenceMap = Map.Make(struct type t = NReference.reference let compare = NReference.compare end)
272 type db = (typ_context * typ option) ReferenceMap.t
274 class type g_status =
276 method extraction_db: db
279 class virtual status =
282 val extraction_db = ReferenceMap.empty
283 method extraction_db = extraction_db
284 method set_extraction_db v = {< extraction_db = v >}
285 method set_extraction_status
286 : 'status. #g_status as 'status -> 'self
287 = fun o -> {< extraction_db = o#extraction_db >}
290 let rec split_kind_prods context =
292 | KArrow (so,ta)-> split_kind_prods (Some ("_",so)::context) ta
293 | KSkip ta -> split_kind_prods (None::context) ta
294 | Type -> context,Type
297 let rec split_typ_prods context =
299 | Arrow (so,ta)-> split_typ_prods (Some ("_",`OfType so)::context) ta
300 | Forall (name,so,ta)-> split_typ_prods (Some (name,`OfKind so)::context) ta
301 | TSkip ta -> split_typ_prods (None::context) ta
302 | _ as t -> context,t
305 let rec glue_ctx_typ ctx typ =
308 | Some (_,`OfType so)::ctx -> glue_ctx_typ ctx (Arrow (so,typ))
309 | Some (name,`OfKind so)::ctx -> glue_ctx_typ ctx (Forall (name,so,typ))
310 | None::ctx -> glue_ctx_typ ctx (TSkip typ)
313 let rec split_typ_lambdas status n ~metasenv context typ =
314 if n = 0 then context,typ
316 match NCicReduction.whd status ~delta:max_int ~subst:[] context typ with
317 | NCic.Lambda (name,s,t) ->
318 split_typ_lambdas status (n-1) ~metasenv ((name,NCic.Decl s)::context) t
320 (* eta-expansion required *)
321 let ty = NCicTypeChecker.typeof status ~metasenv ~subst:[] context t in
322 match NCicReduction.whd status ~delta:max_int ~subst:[] context ty with
323 | NCic.Prod (name,typ,_) ->
324 split_typ_lambdas status (n-1) ~metasenv
325 ((name,NCic.Decl typ)::context)
326 (NCicUntrusted.mk_appl t [NCic.Rel 1])
327 | _ -> assert false (* IMPOSSIBLE *)
331 let context_of_typformer status ~metasenv context =
333 NCic.Const (NReference.Ref (_,NReference.Ind _) as ref)
334 | NCic.Const (NReference.Ref (_,NReference.Def _) as ref)
335 | NCic.Const (NReference.Ref (_,NReference.Decl) as ref)
336 | NCic.Const (NReference.Ref (_,NReference.Fix _) as ref) ->
337 (try fst (ReferenceMap.find ref status#extraction_db)
340 prerr_endline ("REF NOT FOUND: " ^ NReference.string_of_reference ref);
341 assert false (* IMPOSSIBLE *))
342 | NCic.Match _ -> assert false (* TODO ???? *)
345 match List.nth context (n-1) with
346 _,NCic.Decl typ -> typ
347 | _,NCic.Def _ -> assert false (* IMPOSSIBLE *) in
348 let typ_ctx = snd (HExtlib.split_nth n context) in
349 let typ = kind_of status ~metasenv typ_ctx typ in
350 fst (split_kind_prods [] typ)
351 | NCic.Meta _ -> assert false (* TODO *)
352 | NCic.Const (NReference.Ref (_,NReference.Con _))
353 | NCic.Const (NReference.Ref (_,NReference.CoFix _))
354 | NCic.Sort _ | NCic.Implicit _ | NCic.Lambda _ | NCic.LetIn _
355 | NCic.Appl _ | NCic.Prod _ ->
356 assert false (* IMPOSSIBLE *)
358 let rec typ_of status ~metasenv context k =
359 match NCicReduction.whd status ~delta:max_int ~subst:[] context k with
360 | NCic.Prod (b,s,t) ->
361 (* CSC: non-invariant assumed here about "_" *)
362 (match classify status ~metasenv context s with
364 Forall (b, kind_of status ~metasenv context s,
365 typ_of ~metasenv status ((b,NCic.Decl s)::context) t)
367 | `KindOrType -> (* ??? *)
368 Arrow (typ_of status ~metasenv context s,
369 typ_of status ~metasenv ((b,NCic.Decl s)::context) t)
372 TSkip (typ_of status ~metasenv ((b,NCic.Decl s)::context) t)
373 | `Term _ -> assert false (* IMPOSSIBLE *))
376 | NCic.LetIn _ -> assert false (* IMPOSSIBLE *)
377 | NCic.Lambda _ -> assert false (* LAMBDA-LIFT INNER DECLARATION *)
378 | NCic.Rel n -> Var n
379 | NCic.Const ref -> TConst ref
380 | NCic.Appl (he::args) ->
381 let he_context = context_of_typformer status ~metasenv context he in
382 TAppl (typ_of status ~metasenv context he ::
384 (function None -> Unit | Some ty -> typ_of status ~metasenv context ty)
385 (skip_args status ~metasenv context (List.rev he_context,args)))
386 | NCic.Appl _ -> assert false (* TODO: when head is a match/let rec;
387 otherwise NOT A TYPE *)
389 | NCic.Match (_,_,_,_) -> assert false (* TODO *)
392 let rec fomega_subst k t1 =
396 else if n < k then Var n
399 | TConst ref -> TConst ref
401 | Arrow (ty1,ty2) -> Arrow (fomega_subst k t1 ty1, fomega_subst (k+1) t1 ty2)
402 | TSkip t -> TSkip (fomega_subst (k+1) t1 t)
403 | Forall (n,kind,t) -> Forall (n,kind,fomega_subst (k+1) t1 t)
404 | TAppl args -> TAppl (List.map (fomega_subst k t1) args)
406 let fomega_lookup status ref = snd (ReferenceMap.find ref status#extraction_db)
408 let rec fomega_whd status ty =
411 (match fomega_lookup status r with
413 | Some ty -> fomega_whd status ty)
414 | TAppl (TConst r::args) ->
415 (match fomega_lookup status r with
417 | Some ty -> fomega_whd status (List.fold_right (fomega_subst 1) args ty))
420 let rec term_of status ~metasenv context =
424 | NCic.Prod _ -> assert false (* IMPOSSIBLE *)
425 | NCic.Lambda (b,ty,bo) ->
426 (* CSC: non-invariant assumed here about "_" *)
427 (match classify status ~metasenv context ty with
429 TLambda (b,term_of status ~metasenv ((b,NCic.Decl ty)::context) bo)
430 | `KindOrType (* ??? *)
432 Lambda (b, term_of status ~metasenv ((b,NCic.Decl ty)::context) bo)
436 Skip (term_of status ~metasenv ((b,NCic.Decl ty)::context) bo)
437 | `Term _ -> assert false (* IMPOSSIBLE *))
438 | NCic.LetIn (b,ty,t,bo) ->
439 (match classify status ~metasenv context t with
440 | `Term `TypeFormerOrTerm (* ???? *)
442 LetIn (b,term_of status ~metasenv context t,
443 term_of status ~metasenv ((b,NCic.Def (t,ty))::context) bo)
452 (* not in programming languages, we expand it *)
453 term_of status ~metasenv context
454 (NCicSubstitution.subst status ~avoid_beta_redexes:true t bo))
455 | NCic.Rel n -> Rel n
456 | NCic.Const ref -> Const ref
457 | NCic.Appl (he::args) ->
458 eat_args status metasenv
459 (term_of status ~metasenv context he) context
460 (*BUG: recomputing every time the type of the head*)
461 (typ_of status ~metasenv context
462 (NCicTypeChecker.typeof status ~metasenv ~subst:[] context he))
464 | NCic.Appl _ -> assert false (* TODO: when head is a match/let rec;
465 otherwise NOT A TYPE *)
466 | NCic.Meta _ -> assert false (* TODO *)
467 | NCic.Match (ref,_,t,pl) ->
469 let constructors, leftno =
470 let _,leftno,tys,_,n = NCicEnvironment.get_checked_indtys status ref in
471 let _,_,_,cl = List.nth tys n in
474 let rec eat_branch n ty context ctx pat =
478 | Arrow (_, t), _ when n > 0 ->
479 eat_branch (pred n) t context ctx pat
480 | _, _ when n > 0 -> assert false (*BUG: is Top possible here?*)
481 (*CSC: unify the three cases below? *)
482 | Arrow (_, t), NCic.Lambda (name, ty, t') ->
484 (Some (name,`OfType (typ_of status ~metasenv context ty)))::ctx in
485 let context = (name,NCic.Decl ty)::context in
486 eat_branch 0 t context ctx t'
487 | Forall (_,_,t),NCic.Lambda (name, ty, t') ->
489 (Some (name,`OfKind (kind_of status ~metasenv context ty)))::ctx in
490 let context = (name,NCic.Decl ty)::context in
491 eat_branch 0 t context ctx t'
492 | TSkip t,NCic.Lambda (name, ty, t') ->
493 let ctx = None::ctx in
494 let context = (name,NCic.Decl ty)::context in
495 eat_branch 0 t context ctx t'
496 | Top,_ -> assert false (*TODO: HOW??*)
497 (*BUG here, eta-expand!*)
498 | _, _ -> context,ctx, pat
502 (fun (_, name, ty) pat ->
503 (*BUG: recomputing every time the type of the constructor*)
504 let ty = typ_of status ~metasenv context ty in
505 let context,lhs,rhs = eat_branch leftno ty context [] pat in
507 (* UnsafeCoerce not always required *)
508 UnsafeCoerce (term_of status ~metasenv context rhs)
512 with Invalid_argument _ -> assert false
514 (match classify_not_term status [] (NCic.Const ref) with
517 | `Kind -> assert false (* IMPOSSIBLE *)
519 (match patterns_of pl with
521 | [_lhs,rhs] -> rhs (*BUG HERE: Rels are not ok, bound in the _lhs*)
524 Match (ref,term_of status ~metasenv context t, patterns_of pl))
525 and eat_args status metasenv acc context tyhe =
529 match fomega_whd status tyhe with
534 | _ -> term_of status ~metasenv context arg in
535 eat_args status metasenv (mk_appl acc arg) context t tl
537 eat_args status metasenv (Inst acc)
538 context (fomega_subst 1 (typ_of status ~metasenv context arg) t) tl
540 eat_args status metasenv acc context t tl
541 | Top -> assert false (*TODO: HOW??*)
542 | Unit | Var _ | TConst _ | TAppl _ -> assert false (* NOT A PRODUCT *)
552 let object_of_constant status ~metasenv ref bo ty =
553 match classify status ~metasenv [] ty with
555 let ty = kind_of status ~metasenv [] ty in
556 let ctx0,res = split_kind_prods [] ty in
558 split_typ_lambdas status ~metasenv (List.length ctx0) [] bo in
559 (match classify status ~metasenv ctx bo with
561 | `KindOrType -> (* ?? no kind formers in System F_omega *)
565 HExtlib.map_option (fun (_,k) ->
566 (*CSC: BUG here, clashes*)
567 String.uncapitalize (fst n),k) p1)
570 let bo = typ_of status ~metasenv ctx bo in
571 status#set_extraction_db
572 (ReferenceMap.add ref (nicectx,Some bo)
573 status#extraction_db),
574 Success (ref,TypeDefinition((nicectx,res),bo))
575 | `Kind -> status, Erased (* DPM: but not really, more a failure! *)
577 | `Proposition -> status, Erased
578 | `Term _ -> status, Failure "Body of type lambda classified as a term. This is a bug.")
580 | `Proposition -> status, Erased
581 | `KindOrType (* ??? *)
583 (* CSC: TO BE FINISHED, REF NON REGISTERED *)
584 let ty = typ_of status ~metasenv [] ty in
586 Success (ref, TermDefinition (split_typ_prods [] ty, term_of status ~metasenv [] bo))
587 | `Term _ -> status, Failure "Non-term classified as a term. This is a bug."
590 let object_of_inductive status ~metasenv uri ind leftno il =
591 let status,_,rev_tyl =
593 (fun (status,i,res) (_,_,arity,cl) ->
594 match classify_not_term status [] arity with
597 | `Type -> assert false (* IMPOSSIBLE *)
598 | `PropKind -> status,i+1,res
600 let arity = kind_of status ~metasenv [] arity in
601 let ctx,_ = split_kind_prods [] arity in
602 let right,left = HExtlib.split_nth (List.length ctx - leftno) ctx in
604 NReference.reference_of_spec uri (NReference.Ind (ind,i,leftno)) in
606 status#set_extraction_db
607 (ReferenceMap.add ref (ctx,None) status#extraction_db) in
612 NCicReduction.split_prods status ~subst:[] [] leftno ty in
613 let ty = typ_of status ~metasenv ctx ty in
614 NReference.mk_constructor (j+1) ref,ty
617 status,i+1,(ref,left,right,cl)::res
622 | _ -> status, Success (dummyref, Algebraic (List.rev rev_tyl))
625 let object_of status (uri,height,metasenv,subst,obj_kind) =
626 let obj_kind = apply_subst subst obj_kind in
628 | NCic.Constant (_,_,None,ty,_) ->
629 let ref = NReference.reference_of_spec uri NReference.Decl in
630 (match classify status ~metasenv [] ty with
632 let ty = kind_of status ~metasenv [] ty in
633 let ctx,_ as res = split_kind_prods [] ty in
634 status#set_extraction_db
635 (ReferenceMap.add ref (ctx,None) status#extraction_db),
636 Success (ref, TypeDeclaration res)
638 | `Proposition -> status, Erased
640 | `KindOrType (*???*) ->
641 let ty = typ_of status ~metasenv [] ty in
642 status, Success (ref, TermDeclaration (split_typ_prods [] ty))
643 | `Term _ -> status, Failure "Type classified as a term. This is a bug.")
644 | NCic.Constant (_,_,Some bo,ty,_) ->
645 let ref = NReference.reference_of_spec uri (NReference.Def height) in
646 object_of_constant status ~metasenv ref bo ty
647 | NCic.Fixpoint (fix_or_cofix,fs,_) ->
650 (fun (_,_name,recno,ty,bo) (i,status,res) ->
653 NReference.reference_of_spec
654 uri (NReference.Fix (i,recno,height))
656 NReference.reference_of_spec uri (NReference.CoFix i)
658 let status,obj = object_of_constant ~metasenv status ref bo ty in
660 | Success (ref,obj) -> i+1,status, (ref,obj)::res
661 | _ -> i+1,status, res) fs (0,status,[])
663 status, Success (dummyref,LetRec objs)
664 | NCic.Inductive (ind,leftno,il,_) ->
665 object_of_inductive status ~metasenv uri ind leftno il
667 (************************ HASKELL *************************)
669 (* -----------------------------------------------------------------------------
670 * Helper functions I can't seem to find anywhere in the OCaml stdlib?
671 * -----------------------------------------------------------------------------
681 let uncurry f (x, y) =
685 let rec char_list_of_string s =
686 let l = String.length s in
687 let rec aux buffer s =
690 | m -> aux (s.[m - 1]::buffer) s (m - 1)
695 let string_of_char_list s =
699 | x::xs -> aux (String.make 1 x ^ buffer) xs
704 (* ----------------------------------------------------------------------------
705 * Haskell name management: prettyfying valid and idiomatic Haskell identifiers
706 * and type variable names.
707 * ----------------------------------------------------------------------------
710 let remove_underscores_and_mark =
711 let rec aux char_list_buffer positions_buffer position =
713 | [] -> (string_of_char_list char_list_buffer, positions_buffer)
716 aux char_list_buffer (position::positions_buffer) position xs
718 aux (x::char_list_buffer) positions_buffer (position + 1) xs
723 let rec capitalize_marked_positions s =
727 if x < String.length s then
728 let c = Char.uppercase (String.get s x) in
729 let _ = String.set s x c in
730 capitalize_marked_positions s xs
732 capitalize_marked_positions s xs
735 let contract_underscores_and_capitalise =
736 char_list_of_string |>
737 remove_underscores_and_mark |>
738 uncurry capitalize_marked_positions
741 let idiomatic_haskell_type_name_of_string =
742 contract_underscores_and_capitalise |>
746 let idiomatic_haskell_term_name_of_string =
747 contract_underscores_and_capitalise |>
751 (*CSC: code to be changed soon when we implement constructors and
752 we fix the code for term application *)
753 let classify_reference status ref =
754 if ReferenceMap.mem ref status#extraction_db then
757 let NReference.Ref (_,ref) = ref in
759 NReference.Con _ -> `Constructor
760 | NReference.Ind _ -> assert false
764 let capitalize classification name =
765 match classification with
767 | `TypeName -> idiomatic_haskell_type_name_of_string name
770 | `FunctionName -> idiomatic_haskell_term_name_of_string name
773 let pp_ref status ref =
774 capitalize (classify_reference status ref)
775 (NCicPp.r2s status true ref)
777 (* cons avoid duplicates *)
778 let rec (@:::) name l =
779 if name <> "" (* propositional things *) && name.[0] = '_' then
780 let name = String.sub name 1 (String.length name - 1) in
781 let name = if name = "" then "a" else name in
783 else if List.mem name l then (name ^ "'") @::: l
787 let (@::) x l = let x,l = x @::: l in x::l;;
789 let rec pretty_print_type status ctxt =
791 | Var n -> List.nth ctxt (n-1)
793 | Top -> assert false (* ??? *)
794 | TConst ref -> pp_ref status ref
796 bracket size_of_type (pretty_print_type status ctxt) t1 ^ " -> " ^
797 pretty_print_type status ("_"::ctxt) t2
798 | TSkip t -> pretty_print_type status ("_"::ctxt) t
799 | Forall (name, kind, t) ->
800 (*CSC: BUG HERE: avoid clashes due to uncapitalisation*)
801 let name = capitalize `TypeVariable name in
802 let name,ctxt = name@:::ctxt in
803 if size_of_kind kind > 1 then
804 "forall (" ^ name ^ " :: " ^ pretty_print_kind kind ^ "). " ^ pretty_print_type status (name::ctxt) t
806 "forall " ^ name ^ ". " ^ pretty_print_type status (name::ctxt) t
807 | TAppl tl -> String.concat " " (List.map (pretty_print_type status ctxt) tl)
809 let pretty_print_term_context status ctx1 ctx2 =
810 let name_context, rev_res =
812 (fun el (ctx1,rev_res) ->
814 None -> ""@::ctx1,rev_res
815 | Some (name,`OfKind _) -> name@::ctx1,rev_res
816 | Some (name,`OfType typ) ->
817 let name,ctx1 = name@:::ctx1 in
819 ("(" ^ name ^ " :: " ^ pretty_print_type status ctx1 typ ^ ")")::rev_res
821 name_context, String.concat " " (List.rev rev_res)
823 let rec pretty_print_term status ctxt =
825 | Rel n -> List.nth ctxt (n-1)
827 | Const ref -> pp_ref status ref
829 let name = capitalize `BoundVariable name in
830 let name,ctxt = name@:::ctxt in
831 "\\" ^ name ^ " -> " ^ pretty_print_term status (name::ctxt) t
832 | Appl tl -> String.concat " " (List.map (bracket size_of_term (pretty_print_term status ctxt)) tl)
833 | LetIn (name,s,t) ->
834 let name = capitalize `BoundVariable name in
835 let name,ctxt = name@:::ctxt in
836 "let " ^ name ^ " = " ^ pretty_print_term status ctxt s ^ " in " ^
837 pretty_print_term status (name::ctxt) t
839 "error \"Unreachable code\""
841 "unsafeCoerce " ^ bracket size_of_term (pretty_print_term status ctxt) t
842 | Match (r,matched,pl) ->
844 "error \"Case analysis over empty type\""
846 "case " ^ pretty_print_term status ctxt matched ^ " of\n" ^
849 (fun (bound_names,rhs) i ->
850 let ref = NReference.mk_constructor (i+1) r in
851 let name = pp_ref status ref in
852 let ctxt,bound_names =
853 pretty_print_term_context status ctxt bound_names in
855 pretty_print_term status ctxt rhs
857 " " ^ name ^ " " ^ bound_names ^ " -> " ^ body
859 | Skip t -> pretty_print_term status ("[[skipped]]"@::ctxt) t
860 | TLambda (name,t) ->
861 let name = capitalize `TypeVariable name in
862 pretty_print_term status (name@::ctxt) t
863 | Inst t -> pretty_print_term status ctxt t
866 let rec pp_obj status (ref,obj_kind) =
867 let pretty_print_context ctx =
868 String.concat " " (List.rev (snd
870 (fun (x,kind) (l,res) ->
872 if size_of_kind kind > 1 then
873 x::l,("(" ^ x ^ " :: " ^ pretty_print_kind kind ^ ")")::res
876 (HExtlib.filter_map (fun x -> x) ctx) ([],[]))))
878 let namectx_of_ctx ctx =
879 List.fold_right (@::)
880 (List.map (function None -> "" | Some (x,_) -> x) ctx) [] in
882 TypeDeclaration (ctx,_) ->
883 (* data?? unsure semantics: inductive type without constructor, but
884 not matchable apparently *)
885 if List.length ctx = 0 then
886 "data " ^ pp_ref status ref
888 "data " ^ pp_ref status ref ^ " " ^ pretty_print_context ctx
889 | TypeDefinition ((ctx, _),ty) ->
890 let namectx = namectx_of_ctx ctx in
891 if List.length ctx = 0 then
892 "type " ^ pp_ref status ref ^ " = " ^ pretty_print_type status namectx ty
894 "type " ^ pp_ref status ref ^ " " ^ pretty_print_context ctx ^ " = " ^ pretty_print_type status namectx ty
895 | TermDeclaration (ctx,ty) ->
896 let name = pp_ref status ref in
897 name ^ " :: " ^ pretty_print_type status [] (glue_ctx_typ ctx ty) ^ "\n" ^
898 name ^ " = error \"The declaration `" ^ name ^ "' has yet to be defined.\""
899 | TermDefinition ((ctx,ty),bo) ->
900 let name = pp_ref status ref in
901 let namectx = namectx_of_ctx ctx in
903 name ^ " :: " ^ pretty_print_type status namectx (glue_ctx_typ ctx ty) ^ "\n" ^
904 name ^ " = " ^ pretty_print_term status namectx bo
906 String.concat "\n" (List.map (pp_obj status) l)
910 (fun ref,left,right,cl ->
911 "data " ^ pp_ref status ref ^ " " ^
912 pretty_print_context (right@left) ^ " where\n " ^
913 String.concat "\n " (List.map
915 let namectx = namectx_of_ctx left in
916 pp_ref status ref ^ " :: " ^
917 pretty_print_type status namectx tys
920 (* inductive and records missing *)
922 let haskell_of_obj status (uri,_,_,_,_ as obj) =
923 let status, obj = object_of status obj in
926 Erased -> "-- [?] " ^ NUri.name_of_uri uri ^ " erased due to term being propositionally irrelevant.\n"
927 | OutsideTheory -> "-- [?] " ^ NUri.name_of_uri uri ^ " erased due to image of term under extraction residing outside Fω.\n"
928 | Failure msg -> "-- [?] " ^ NUri.name_of_uri uri ^ " FAILURE: " ^ msg ^ "\n"
929 | Success o -> pp_obj status o ^ "\n"