(***********************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* mt () | [i] -> str"Position [" ++ int i ++ str"] is implicit" | l -> str"Positions [" ++ prlist_with_sep (fun () -> str "; ") int l ++ str"] are implicit" let print_impl_args l = print_impl_args_by_pos (positions_of_implicits l) (* To be improved; the type should be used to provide the types in the abstractions. This should be done recursively inside prterm, so that the pretty-print of a proposition (P:(nat->nat)->Prop)(P [u]u) synthesizes the type nat of the abstraction on u *) let print_named_def name body typ = let pbody = prterm body in let ptyp = prtype typ in (str "*** [" ++ str name ++ str " " ++ hov 0 (str ":=" ++ brk (1,2) ++ pbody ++ spc () ++ str ":" ++ brk (1,2) ++ ptyp) ++ str "]" ++ fnl ()) let print_named_assum name typ = (str "*** [" ++ str name ++ str " : " ++ prtype typ ++ str "]" ++ fnl ()) let print_named_decl (id,c,typ) = let s = string_of_id id in match c with | Some body -> print_named_def s body typ | None -> print_named_assum s typ let assumptions_for_print lna = List.fold_right (fun na env -> add_name na env) lna empty_names_context let implicit_args_id id l = if List.exists is_status_implicit l then (str"For " ++ pr_id id ++ str": " ++ print_impl_args l ++ fnl ()) else (mt ()) let implicit_args_msg sp mipv = (prvecti (fun i mip -> let imps = implicits_of_global (IndRef (sp,i)) in ((implicit_args_id mip.mind_typename imps) ++ prvecti (fun j idc -> let imps = implicits_of_global (ConstructRef ((sp,i),j+1)) in (implicit_args_id idc imps)) mip.mind_consnames )) mipv) let print_params env params = if List.length params = 0 then (mt ()) else (str "[" ++ pr_rel_context env params ++ str "]" ++ brk(1,2)) let rec contract_types = function | [] -> [] | (id,c)::l -> match contract_types l with | (idl,c')::l' when eq_constr c c' -> (id::idl,c')::l' | l' -> ([id],c)::l' let print_constructors envpar names types = let pc = prlist_with_sep (fun () -> brk(1,0) ++ str "| ") (fun (id,c) -> prlist_with_sep pr_coma pr_id id ++ str " : " ++ prterm_env envpar c) (contract_types (Array.to_list (array_map2 (fun n t -> (n,t)) names types))) in hv 0 (str " " ++ pc) let build_inductive sp tyi = let (mib,mip) = Global.lookup_inductive (sp,tyi) in let params = mip.mind_params_ctxt in let args = extended_rel_list 0 params in let env = Global.env() in let arity = hnf_prod_applist env mip.mind_user_arity args in let cstrtypes = arities_of_constructors env (sp,tyi) in let cstrtypes = Array.map (fun c -> hnf_prod_applist env c args) cstrtypes in let cstrnames = mip.mind_consnames in (IndRef (sp,tyi), params, arity, cstrnames, cstrtypes) let print_one_inductive (sp,tyi) = let (ref, params, arity, cstrnames, cstrtypes) = build_inductive sp tyi in let env = Global.env () in let envpar = push_rel_context params env in hov 0 ( pr_global (IndRef (sp,tyi)) ++ brk(1,2) ++ print_params env params ++ str ": " ++ prterm_env envpar arity ++ str " :=") ++ brk(0,2) ++ print_constructors envpar cstrnames cstrtypes let pr_mutual_inductive finite indl = hov 0 ( str (if finite then "Inductive " else "CoInductive ") ++ prlist_with_sep (fun () -> fnl () ++ str" with ") print_one_inductive indl) ++ fnl () let print_mutual sp = let (mib,mip) = Global.lookup_inductive (sp,0) in let mipv = mib.mind_packets in let names = list_tabulate (fun x -> (sp,x)) (Array.length mipv) in pr_mutual_inductive mib.mind_finite names ++ implicit_args_msg sp mipv let print_section_variable sp = let d = get_variable sp in let l = implicits_of_global (VarRef sp) in (print_named_decl d ++ print_impl_args l) let print_body = function | Some lc -> prterm (Declarations.force lc) | None -> (str"") let print_typed_body (val_0,typ) = (print_body val_0 ++ fnl () ++ str " : " ++ prtype typ ++ fnl ()) let print_constant with_values sep sp = let cb = Global.lookup_constant sp in let val_0 = cb.const_body in let typ = cb.const_type in let impls = implicits_of_global (ConstRef sp) in hov 0 ((match val_0 with | None -> (str"*** [ " ++ print_basename sp ++ str " : " ++ cut () ++ prtype typ ++ str" ]" ++ fnl ()) | _ -> (print_basename sp ++ str sep ++ cut () ++ if with_values then print_typed_body (val_0,typ) else (prtype typ ++ fnl ()))) ++ print_impl_args impls) let print_inductive sp = (print_mutual sp) let print_syntactic_def sep kn = let qid = Nametab.shortest_qualid_of_syndef kn in let c = Syntax_def.search_syntactic_definition dummy_loc kn in (str" Syntactic Definition " ++ pr_qualid qid ++ str sep ++ Constrextern.without_symbols pr_rawterm c ++ fnl ()) let print_leaf_entry with_values sep ((sp,kn as oname),lobj) = let tag = object_tag lobj in match (oname,tag) with | (_,"VARIABLE") -> print_section_variable (basename sp) ++ fnl () | (_,"CONSTANT") -> print_constant with_values sep kn ++ fnl () | (_,"INDUCTIVE") -> print_inductive kn ++ fnl () | (_,"MODULE") -> let (mp,_,l) = repr_kn kn in print_module with_values (MPdot (mp,l)) ++ fnl () | (_,"MODULE TYPE") -> print_modtype kn ++ fnl () | (_,"AUTOHINT") -> (* (str" Hint Marker" ++ fnl ())*) (mt ()) | (_,"GRAMMAR") -> (* (str" Grammar Marker" ++ fnl ())*) (mt ()) | (_,"SYNTAXCONSTANT") -> print_syntactic_def sep kn ++ fnl () | (_,"PPSYNTAX") -> (* (str" Syntax Marker" ++ fnl ())*) (mt ()) | (_,"TOKEN") -> (* (str" Token Marker" ++ fnl ())*) (mt ()) | (_,"CLASS") -> (* (str" Class Marker" ++ fnl ())*) (mt ()) | (_,"COERCION") -> (* (str" Coercion Marker" ++ fnl ())*) (mt ()) | (_,"REQUIRE") -> (* (str" Require Marker" ++ fnl ())*) (mt ()) | (_,"END-SECTION") -> (mt ()) | (_,"STRUCTURE") -> (mt ()) (* To deal with forgotten cases... *) | (_,s) -> (mt ()) (* | (_,s) -> (str(string_of_path sp) ++ str" : " ++ str"Unrecognized object " ++ str s ++ fnl ()) *) let rec print_library_entry with_values ent = let sep = if with_values then " = " else " : " in let pr_name (sp,_) = pr_id (basename sp) in match ent with | (oname,Lib.Leaf lobj) -> (print_leaf_entry with_values sep (oname,lobj)) | (oname,Lib.OpenedSection (dir,_)) -> (str " >>>>>>> Section " ++ pr_name oname ++ fnl ()) | (oname,Lib.ClosedSection _) -> (str " >>>>>>> Closed Section " ++ pr_name oname ++ fnl ()) | (_,Lib.CompilingLibrary (dir,_)) -> (str " >>>>>>> Library " ++ pr_dirpath dir ++ fnl ()) | (oname,Lib.OpenedModule _) -> (str " >>>>>>> Module " ++ pr_name oname ++ fnl ()) | (oname,Lib.OpenedModtype _) -> (str " >>>>>>> Module Type " ++ pr_name oname ++ fnl ()) | (_,Lib.FrozenState _) -> (mt ()) and print_context with_values = let rec prec = function | h::rest -> (prec rest ++ print_library_entry with_values h) | [] -> (mt ()) in prec let print_full_context () = print_context true (Lib.contents_after None) let print_full_context_typ () = print_context false (Lib.contents_after None) (* For printing an inductive definition with its constructors and elimination, assume that the declaration of constructors and eliminations follows the definition of the inductive type *) let list_filter_vec f vec = let rec frec n lf = if n < 0 then lf else if f vec.(n) then frec (n-1) (vec.(n)::lf) else frec (n-1) lf in frec (Array.length vec -1) [] (* This is designed to print the contents of an opened section *) let read_sec_context r = let loc,qid = qualid_of_reference r in let dir = try Nametab.locate_section qid with Not_found -> user_err_loc (loc,"read_sec_context", str "Unknown section") in let rec get_cxt in_cxt = function | ((_,Lib.OpenedSection ((dir',_),_)) as hd)::rest -> if dir = dir' then (hd::in_cxt) else get_cxt (hd::in_cxt) rest | ((_,Lib.ClosedSection (_,_,ctxt)) as hd)::rest -> error "Cannot print the contents of a closed section" | [] -> [] | hd::rest -> get_cxt (hd::in_cxt) rest in let cxt = (Lib.contents_after None) in List.rev (get_cxt [] cxt) let print_sec_context sec = print_context true (read_sec_context sec) let print_sec_context_typ sec = print_context false (read_sec_context sec) let print_judgment env {uj_val=trm;uj_type=typ} = print_typed_value_in_env env (trm, typ) let print_safe_judgment env j = let trm = Safe_typing.j_val j in let typ = Safe_typing.j_type j in print_typed_value_in_env env (trm, typ) let print_eval red_fun env {uj_val=trm;uj_type=typ} = let ntrm = red_fun env Evd.empty trm in (str " = " ++ print_judgment env {uj_val = ntrm; uj_type = typ}) let print_name r = let loc,qid = qualid_of_reference r in try let sp = Nametab.locate_obj qid in let (oname,lobj) = let (oname,entry) = List.find (fun en -> (fst (fst en)) = sp) (Lib.contents_after None) in match entry with | Lib.Leaf obj -> (oname,obj) | _ -> raise Not_found in print_leaf_entry true " = " (oname,lobj) with Not_found -> try match Nametab.locate qid with | ConstRef sp -> print_constant true " = " sp | IndRef (sp,_) -> print_inductive sp | ConstructRef ((sp,_),_) -> print_inductive sp | VarRef sp -> print_section_variable sp with Not_found -> try (* Var locale de but, pas var de section... donc pas d'implicits *) let dir,str = repr_qualid qid in if (repr_dirpath dir) <> [] then raise Not_found; let (_,c,typ) = Global.lookup_named str in (print_named_decl (str,c,typ)) with Not_found -> try let kn = Nametab.locate_syntactic_definition qid in print_syntactic_def " = " kn with Not_found -> try let globdir = Nametab.locate_dir qid in match globdir with DirModule (dirpath,(mp,_)) -> print_module (printable_body dirpath) mp | _ -> raise Not_found with Not_found -> try print_modtype (Nametab.locate_modtype qid) with Not_found -> user_err_loc (loc,"print_name",pr_qualid qid ++ spc () ++ str "not a defined object") let print_opaque_name qid = let sigma = Evd.empty in let env = Global.env () in let sign = Global.named_context () in match global qid with | ConstRef cst -> let cb = Global.lookup_constant cst in if cb.const_body <> None then print_constant true " = " cst else error "not a defined constant" | IndRef (sp,_) -> print_mutual sp | ConstructRef cstr -> let ty = Inductive.type_of_constructor env cstr in print_typed_value (mkConstruct cstr, ty) | VarRef id -> let (_,c,ty) = lookup_named id env in print_named_decl (id,c,ty) let print_local_context () = let env = Lib.contents_after None in let rec print_var_rec = function | [] -> (mt ()) | (oname,Lib.Leaf lobj)::rest -> if "VARIABLE" = object_tag lobj then let d = get_variable (basename (fst oname)) in (print_var_rec rest ++ print_named_decl d) else print_var_rec rest | _::rest -> print_var_rec rest and print_last_const = function | (oname,Lib.Leaf lobj)::rest -> (match object_tag lobj with | "CONSTANT" -> let kn = snd oname in let {const_body=val_0;const_type=typ} = Global.lookup_constant kn in (print_last_const rest ++ print_basename kn ++str" = " ++ print_typed_body (val_0,typ)) | "INDUCTIVE" -> let kn = snd oname in (print_last_const rest ++print_mutual kn ++ fnl ()) | "VARIABLE" -> (mt ()) | _ -> print_last_const rest) | _ -> (mt ()) in (print_var_rec env ++ print_last_const env) let unfold_head_fconst = let rec unfrec k = match kind_of_term k with | Const cst -> constant_value (Global.env ()) cst | Lambda (na,t,b) -> mkLambda (na,t,unfrec b) | App (f,v) -> appvect (unfrec f,v) | _ -> k in unfrec (* for debug *) let inspect depth = let rec inspectrec n res env = if n=0 or env=[] then res else inspectrec (n-1) (List.hd env::res) (List.tl env) in let items = List.rev (inspectrec depth [] (Lib.contents_after None)) in print_context false items (*************************************************************************) (* Pretty-printing functions coming from classops.ml *) open Classops let print_coercion_value v = prterm (get_coercion_value v) let print_index_coercion c = let _,v = coercion_info_from_index c in print_coercion_value v let print_class i = let cl,_ = class_info_from_index i in (str (string_of_class cl)) let print_path ((i,j),p) = (str"[" ++ prlist_with_sep (fun () -> (str"; ")) (fun c -> print_index_coercion c) p ++ str"] : " ++ print_class i ++ str" >-> " ++ print_class j) let _ = Classops.install_path_printer print_path let print_graph () = (prlist_with_sep pr_fnl print_path (inheritance_graph())) let print_classes () = (prlist_with_sep pr_spc (fun (_,(cl,x)) -> (str (string_of_class cl) (* ++ str(string_of_strength x.cl_strength) *))) (classes())) let print_coercions () = (prlist_with_sep pr_spc (fun (_,(_,v)) -> (print_coercion_value v)) (coercions())) let index_of_class cl = try fst (class_info cl) with _ -> errorlabstrm "index_of_class" (str(string_of_class cl) ++ str" is not a defined class") let print_path_between cls clt = let i = index_of_class cls in let j = index_of_class clt in let p = try lookup_path_between (i,j) with _ -> errorlabstrm "index_cl_of_id" (str"No path between " ++str(string_of_class cls) ++ str" and " ++str(string_of_class clt)) in print_path ((i,j),p) (*************************************************************************)